From 4a7a987ba851b3634c0d6805adf55e34fdb4648f Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Fri, 27 Sep 2024 21:36:55 +0000 Subject: [PATCH 001/174] feat: add annotations for autolabel workflow (#17188) improve displayed messages --- scripts/autolabel.lean | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/scripts/autolabel.lean b/scripts/autolabel.lean index ccf7d7321c35e..922e98d1fee24 100644 --- a/scripts/autolabel.lean +++ b/scripts/autolabel.lean @@ -239,12 +239,12 @@ to add the label to the PR. - `2`: invalid labels defined - `3`: ~labels do not cover all of `Mathlib/`~ (unused; only emitting warning) -/ -unsafe def main (args : List String): IO Unit := do +unsafe def main (args : List String): IO UInt32 := do if args.length > 1 then println s!"::error:: autolabel: invalid number of arguments ({args.length}), \ expected at most 1. Please run without arguments or provide the target PR's \ number as a single argument!" - IO.Process.exit 1 + return 1 let prNumber? := args[0]? -- test: validate that all paths in `mathlibLabels` actually exist @@ -267,7 +267,7 @@ unsafe def main (args : List String): IO Unit := do Please update `{ ``AutoLabel.mathlibLabels }`!" valid := false unless valid do - IO.Process.exit 2 + return 2 -- test: validate that the labels cover all of the `Mathlib/` folder let notMatchedPaths ← findUncoveredPaths "Mathlib" (exceptions := mathlibUnlabelled) @@ -279,7 +279,7 @@ unsafe def main (args : List String): IO Unit := do s!"Incomplete `{ ``AutoLabel.mathlibLabels }`" s!"the following paths inside `Mathlib/` are not covered \ by any label: {notMatchedPaths} Please modify `AutoLabel.mathlibLabels` accordingly!" - -- IO.Process.exit 3 + -- return 3 -- get the modified files let gitDiff ← IO.Process.run { @@ -290,21 +290,21 @@ unsafe def main (args : List String): IO Unit := do -- find labels covering the modified files let labels := getMatchingLabels modifiedFiles + println s!"::notice::Applicable labels: {labels}" + match labels with | #[] => - println s!"No applicable labels found!" + println s!"::warning::no label to add" | #[label] => - println s!"Exactly one label found: {label}" match prNumber? with | some n => let _ ← IO.Process.run { cmd := "gh", args := #["pr", "edit", n, "--add-label", label] } - println s!"Added label: {label}" + println s!"::notice::added label: {label}" | none => - println s!"No PR-number provided, skipping adding labels. \ + println s!"::warning::no PR-number provided, not adding labels. \ (call `lake exe autolabel 150602` to add the labels to PR `150602`)" - | labels => - println s!"Multiple labels found: {labels}" - println s!"Not adding any label." - IO.Process.exit 0 + | _ => + println s!"::notice::not adding multiple labels: {labels}" + return 0 From 6c1ff14431a4b2fccd3fb3120b80172312e355e1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 27 Sep 2024 21:45:34 +0000 Subject: [PATCH 002/174] chore(UniqueProds): fix some `to_additive` names (#17072) --- Mathlib/Algebra/Group/UniqueProds/Basic.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Group/UniqueProds/Basic.lean b/Mathlib/Algebra/Group/UniqueProds/Basic.lean index c243d2138661e..99d252eee8364 100644 --- a/Mathlib/Algebra/Group/UniqueProds/Basic.lean +++ b/Mathlib/Algebra/Group/UniqueProds/Basic.lean @@ -70,13 +70,16 @@ variable {G H : Type*} [Mul G] [Mul H] {A B : Finset G} {a0 b0 : G} theorem of_subsingleton [Subsingleton G] : UniqueMul A B a0 b0 := by simp [UniqueMul, eq_iff_true_of_subsingleton] -@[to_additive] +@[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) : ∃ 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 exact ⟨a, ha, b, hb, fun _ _ ha' hb' _ ↦ ⟨hA1 ha' ha, hB1 hb' hb⟩⟩ +@[deprecated (since := "2024-09-23")] +alias _root_.UniqueAdd.of_card_nonpos := UniqueAdd.of_card_le_one + @[to_additive] theorem mt (h : UniqueMul A B a0 b0) : ∀ ⦃a b⦄, a ∈ A → b ∈ B → a ≠ a0 ∨ b ≠ b0 → a * b ≠ a0 * b0 := fun _ _ ha hb k ↦ by @@ -113,7 +116,7 @@ theorem iff_existsUnique (aA : a0 ∈ A) (bB : b0 ∈ B) : exact Prod.mk.inj_iff.mp (J (x, y) ⟨Finset.mk_mem_product hx hy, l⟩))⟩ open Finset in -@[to_additive] +@[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 simp_rw [card_le_one_iff, mem_filter, mem_product] @@ -124,6 +127,9 @@ theorem iff_card_le_one [DecidableEq G] (ha0 : a0 ∈ A) (hb0 : b0 ∈ B) : · rw [h1.2, h2.2] · exact Prod.ext_iff.1 (@h (a, b) (a0, b0) ⟨⟨ha, hb⟩, he⟩ ⟨⟨ha0, hb0⟩, rfl⟩) +@[deprecated (since := "2024-09-23")] +alias _root_.UniqueAdd.iff_card_nonpos := UniqueAdd.iff_card_le_one + -- Porting note: mathport warning: expanding binder collection -- (ab «expr ∈ » [finset.product/multiset.product/set.prod/list.product](A, B)) -/ @[to_additive] From 2b108d1e9018455f2a4d06b17b06a0f689f42278 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Fri, 27 Sep 2024 22:13:25 +0000 Subject: [PATCH 003/174] feat(FiberCategory/Fibered): define fibered functors (#12982) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the class `IsFibered p`, which captures what it means for a functor `p : 𝒳 ⥤ 𝒮` to be a fibered category, defined as in SGA 1 VI 6.1. We also provide an alternate constructor corresponding to the definition as in [02XM](https://stacks.math.columbia.edu/tag/02XM). Co-authored-by: Paul Lezeau Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../FiberedCategory/Cartesian.lean | 38 +++++- .../FiberedCategory/Fibered.lean | 115 +++++++++++++++++- .../FiberedCategory/HomLift.lean | 2 +- 3 files changed, 147 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean b/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean index c2f4a20e934c0..b6f222bcedcbc 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean @@ -119,6 +119,7 @@ lemma map_self : IsCartesian.map p f φ φ = 𝟙 a := by /-- The canonical isomorphism between the domains of two cartesian morphisms lying over the same object. -/ +@[simps] noncomputable def domainUniqueUpToIso {a' : 𝒳} (φ' : a' ⟶ b) [IsCartesian p f φ'] : a' ≅ a where hom := IsCartesian.map p f φ φ' inv := IsCartesian.map p f φ' φ @@ -131,6 +132,14 @@ noncomputable def domainUniqueUpToIso {a' : 𝒳} (φ' : a' ⟶ b) [IsCartesian apply IsCartesian.ext p (p.map φ) φ simp only [assoc, fac, id_comp] +instance domainUniqueUpToIso_inv_isHomLift {a' : 𝒳} (φ' : a' ⟶ b) [IsCartesian p f φ'] : + IsHomLift p (𝟙 R) (domainUniqueUpToIso p f φ φ').hom := + domainUniqueUpToIso_hom p f φ φ' ▸ IsCartesian.map_isHomLift p f φ φ' + +instance domainUniqueUpToIso_hom_isHomLift {a' : 𝒳} (φ' : a' ⟶ b) [IsCartesian p f φ'] : + IsHomLift p (𝟙 R) (domainUniqueUpToIso p f φ φ').inv := + domainUniqueUpToIso_inv p f φ φ' ▸ IsCartesian.map_isHomLift p f φ' φ + /-- Precomposing a cartesian morphism with an isomorphism lifting the identity is cartesian. -/ instance of_iso_comp {a' : 𝒳} (φ' : a' ≅ a) [IsHomLift p (𝟙 R) φ'.hom] : IsCartesian p f (φ'.hom ≫ φ) where @@ -351,15 +360,34 @@ lemma isIso_of_base_isIso (φ : a ⟶ b) [IsStronglyCartesian p f φ] [IsIso f] end +section + +variable {R R' S : 𝒮} {a a' b : 𝒳} {f : R ⟶ S} {f' : R' ⟶ S} {g : R' ≅ R} + /-- The canonical isomorphism between the domains of two strongly cartesian morphisms lying over isomorphic objects. -/ -noncomputable def domainIsoOfBaseIso {R R' S : 𝒮} {a a' b : 𝒳} {f : R ⟶ S} {f' : R' ⟶ S} - {g : R' ≅ R} (h : f' = g.hom ≫ f) (φ : a ⟶ b) (φ' : a' ⟶ b) [IsStronglyCartesian p f φ] - [IsStronglyCartesian p f' φ'] : a' ≅ a where +@[simps] +noncomputable def domainIsoOfBaseIso (h : f' = g.hom ≫ f) (φ : a ⟶ b) (φ' : a' ⟶ b) + [IsStronglyCartesian p f φ] [IsStronglyCartesian p f' φ'] : a' ≅ a where hom := map p f φ h φ' - inv := by - convert map p f' φ' (congrArg (g.inv ≫ ·) h.symm) φ + inv := + haveI : p.IsHomLift ((fun x ↦ g.inv ≫ x) (g.hom ≫ f)) φ := by + simpa using IsCartesian.toIsHomLift + map p f' φ' (congrArg (g.inv ≫ ·) h.symm) φ + +instance domainUniqueUpToIso_inv_isHomLift (h : f' = g.hom ≫ f) (φ : a ⟶ b) (φ' : a' ⟶ b) + [IsStronglyCartesian p f φ] [IsStronglyCartesian p f' φ'] : + IsHomLift p g.hom (domainIsoOfBaseIso p h φ φ').hom := + domainIsoOfBaseIso_hom p h φ φ' ▸ IsStronglyCartesian.map_isHomLift p f φ h φ' + +instance domainUniqueUpToIso_hom_isHomLift (h : f' = g.hom ≫ f) (φ : a ⟶ b) (φ' : a' ⟶ b) + [IsStronglyCartesian p f φ] [IsStronglyCartesian p f' φ'] : + IsHomLift p g.inv (domainIsoOfBaseIso p h φ φ').inv := by + haveI : p.IsHomLift ((fun x ↦ g.inv ≫ x) (g.hom ≫ f)) φ := by simpa using IsCartesian.toIsHomLift + simpa using IsStronglyCartesian.map_isHomLift p f' φ' (congrArg (g.inv ≫ ·) h.symm) φ + +end end IsStronglyCartesian diff --git a/Mathlib/CategoryTheory/FiberedCategory/Fibered.lean b/Mathlib/CategoryTheory/FiberedCategory/Fibered.lean index c3342d615089e..b826fc429e660 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/Fibered.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/Fibered.lean @@ -17,6 +17,15 @@ This file defines what it means for a functor `p : 𝒳 ⥤ 𝒮` to be (pre)fib - `IsPreFibered p` expresses `𝒳` is fibered over `𝒮` via a functor `p : 𝒳 ⥤ 𝒮`, as in SGA VI.6.1. This means that any morphism in the base `𝒮` can be lifted to a cartesian morphism in `𝒳`. +- `IsFibered p` expresses `𝒳` is fibered over `𝒮` via a functor `p : 𝒳 ⥤ 𝒮`, as in SGA VI.6.1. +This means that it is prefibered, and that the composition of any two cartesian morphisms is +cartesian. + +In the literature one often sees the notion of a fibered category defined as the existence of +strongly cartesian morphisms lying over any given morphism in the base. This is equivalent to the +notion above, and we give an alternate constructor `IsFibered.of_exists_isCartesian'` for +constructing a fibered category this way. + ## Implementation The constructor of `IsPreFibered` is called `exists_isCartesian'`. The reason for the prime is that @@ -47,7 +56,18 @@ protected lemma IsPreFibered.exists_isCartesian (p : 𝒳 ⥤ 𝒮) [p.IsPreFibe (ha : p.obj a = S) (f : R ⟶ S) : ∃ (b : 𝒳) (φ : b ⟶ a), IsCartesian p f φ := by subst ha; exact IsPreFibered.exists_isCartesian' f -namespace IsPreFibered +/-- Definition of a fibered category. + +See SGA 1 VI.6.1. -/ +class Functor.IsFibered (p : 𝒳 ⥤ 𝒮) extends IsPreFibered p : Prop where + comp {R S T : 𝒮} (f : R ⟶ S) (g : S ⟶ T) {a b c : 𝒳} (φ : a ⟶ b) (ψ : b ⟶ c) + [IsCartesian p f φ] [IsCartesian p g ψ] : IsCartesian p (f ≫ g) (φ ≫ ψ) + +instance (p : 𝒳 ⥤ 𝒮) [p.IsFibered] {R S T : 𝒮} (f : R ⟶ S) (g : S ⟶ T) {a b c : 𝒳} (φ : a ⟶ b) + (ψ : b ⟶ c) [IsCartesian p f φ] [IsCartesian p g ψ] : IsCartesian p (f ≫ g) (φ ≫ ψ) := + IsFibered.comp f g φ ψ + +namespace Functor.IsPreFibered open IsCartesian @@ -70,6 +90,97 @@ instance pullbackMap.IsCartesian : IsCartesian p f (pullbackMap ha f) := lemma pullbackObj_proj : p.obj (pullbackObj ha f) = R := domain_eq p f (pullbackMap ha f) -end IsPreFibered +end Functor.IsPreFibered + +namespace Functor.IsFibered + +open IsCartesian IsPreFibered + +/-- In a fibered category, any cartesian morphism is strongly cartesian. -/ +instance isStronglyCartesian_of_isCartesian (p : 𝒳 ⥤ 𝒮) [p.IsFibered] {R S : 𝒮} (f : R ⟶ S) + {a b : 𝒳} (φ : a ⟶ b) [p.IsCartesian f φ] : p.IsStronglyCartesian f φ where + universal_property' g φ' hφ' := by + -- Let `ψ` be a cartesian arrow lying over `g` + let ψ := pullbackMap (domain_eq p f φ) g + -- Let `τ` be the map induced by the universal property of `ψ ≫ φ`. + let τ := IsCartesian.map p (g ≫ f) (ψ ≫ φ) φ' + use τ ≫ ψ + -- It is easily verified that `τ ≫ ψ` lifts `g` and `τ ≫ ψ ≫ φ = φ'` + refine ⟨⟨inferInstance, by simp only [assoc, IsCartesian.fac, τ]⟩, ?_⟩ + -- It remains to check that `τ ≫ ψ` is unique. + -- So fix another lift `π` of `g` satisfying `π ≫ φ = φ'`. + intro π ⟨hπ, hπ_comp⟩ + -- Write `π` as `π = τ' ≫ ψ` for some `τ'` induced by the universal property of `ψ`. + rw [← fac p g ψ π] + -- It remains to show that `τ' = τ`. This follows again from the universal property of `ψ`. + congr 1 + apply map_uniq + rwa [← assoc, IsCartesian.fac] + +/-- In a category which admits strongly cartesian pullbacks, any cartesian morphism is +strongly cartesian. This is a helper-lemma for the fact that admitting strongly cartesian pullbacks +implies being fibered. -/ +lemma isStronglyCartesian_of_exists_isCartesian (p : 𝒳 ⥤ 𝒮) (h : ∀ (a : 𝒳) (R : 𝒮) + (f : R ⟶ p.obj a), ∃ (b : 𝒳) (φ : b ⟶ a), IsStronglyCartesian p f φ) {R S : 𝒮} (f : R ⟶ S) + {a b : 𝒳} (φ : a ⟶ b) [p.IsCartesian f φ] : p.IsStronglyCartesian f φ := by + constructor + intro c g φ' hφ' + subst_hom_lift p f φ; clear a b R S + -- Let `ψ` be a cartesian arrow lying over `g` + obtain ⟨a', ψ, hψ⟩ := h _ _ (p.map φ) + -- Let `τ' : c ⟶ a'` be the map induced by the universal property of `ψ` + let τ' := IsStronglyCartesian.map p (p.map φ) ψ (f':= g ≫ p.map φ) rfl φ' + -- Let `Φ : a' ≅ a` be natural isomorphism induced between `φ` and `ψ`. + let Φ := domainUniqueUpToIso p (p.map φ) φ ψ + -- The map induced by `φ` will be `τ' ≫ Φ.hom` + use τ' ≫ Φ.hom + -- It is easily verified that `τ' ≫ Φ.hom` lifts `g` and `τ' ≫ Φ.hom ≫ φ = φ'` + refine ⟨⟨by simp only [Φ]; infer_instance, ?_⟩, ?_⟩ + · simp [τ', Φ, IsStronglyCartesian.map_uniq p (p.map φ) ψ rfl φ'] + -- It remains to check that it is unique. This follows from the universal property of `ψ`. + intro π ⟨hπ, hπ_comp⟩ + rw [← Iso.comp_inv_eq] + apply IsStronglyCartesian.map_uniq p (p.map φ) ψ rfl φ' + simp [hπ_comp, Φ] + +/-- Alternate constructor for `IsFibered`, a functor `p : 𝒳 ⥤ 𝒴` is fibered if any diagram of the +form +``` + a + - + | + v +R --f--> p(a) +``` +admits a strongly cartesian lift `b ⟶ a` of `f`. -/ +lemma of_exists_isStronglyCartesian {p : 𝒳 ⥤ 𝒮} + (h : ∀ (a : 𝒳) (R : 𝒮) (f : R ⟶ p.obj a), + ∃ (b : 𝒳) (φ : b ⟶ a), IsStronglyCartesian p f φ) : + IsFibered p where + exists_isCartesian' := by + intro a R f + obtain ⟨b, φ, hφ⟩ := h a R f + refine ⟨b, φ, inferInstance⟩ + comp := fun R S T f g {a b c} φ ψ _ _ => + have : p.IsStronglyCartesian f φ := isStronglyCartesian_of_exists_isCartesian p h _ _ + have : p.IsStronglyCartesian g ψ := isStronglyCartesian_of_exists_isCartesian p h _ _ + inferInstance + +/-- Given a diagram +``` + a + - + | + v +T --g--> R --f--> S +``` +we have an isomorphism `T ×_S a ≅ T ×_R (R ×_S a)` -/ +noncomputable def pullbackPullbackIso {p : 𝒳 ⥤ 𝒮} [IsFibered p] + {R S T : 𝒮} {a : 𝒳} (ha : p.obj a = S) (f : R ⟶ S) (g : T ⟶ R) : + pullbackObj ha (g ≫ f) ≅ pullbackObj (pullbackObj_proj ha f) g := + domainUniqueUpToIso p (g ≫ f) (pullbackMap (pullbackObj_proj ha f) g ≫ pullbackMap ha f) + (pullbackMap ha (g ≫ f)) + +end Functor.IsFibered end CategoryTheory diff --git a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean index 595fb845dc473..8de8bf2c92401 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean @@ -142,7 +142,7 @@ instance comp_lift_id_left {a b c : 𝒳} {S T : 𝒮} (f : S ⟶ T) (ψ : b ⟶ lemma comp_lift_id_left' {a b c : 𝒳} (R : 𝒮) (φ : a ⟶ b) [p.IsHomLift (𝟙 R) φ] {S T : 𝒮} (f : S ⟶ T) (ψ : b ⟶ c) [p.IsHomLift f ψ] : p.IsHomLift f (φ ≫ ψ) := by obtain rfl : R = S := by rw [← codomain_eq p (𝟙 R) φ, domain_eq p f ψ] - simpa using inferInstanceAs (p.IsHomLift (𝟙 R ≫ f) (φ ≫ ψ)) + infer_instance lemma eqToHom_domain_lift_id {p : 𝒳 ⥤ 𝒮} {a b : 𝒳} (hab : a = b) {R : 𝒮} (hR : p.obj a = R) : p.IsHomLift (𝟙 R) (eqToHom hab) := by From 8d426f12d4550cf82ee62f169425f78684c746d1 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 28 Sep 2024 07:48:11 +0000 Subject: [PATCH 004/174] feat: the inverse of an analytic partial homeo is also analytic (#17170) We already have the results for the inverse of formal multilinear series, so this PR is mostly just extending our API. However, this showed that our API is suboptimal (notably with respect to the choice of constant coefficients in inverses of formal multilinear series), so we also tweak the definitions of `FormalMultilinearSeries.id` and `FormalMultilinearSeries.leftInv/rightInv` to make them more usable. Co-authored-by: Johan Commelin --- Mathlib/Analysis/Analytic/Basic.lean | 7 + Mathlib/Analysis/Analytic/Composition.lean | 90 +++-- Mathlib/Analysis/Analytic/Inverse.lean | 339 ++++++++++++------ .../Algebra/InfiniteSum/Constructions.lean | 30 +- 4 files changed, 329 insertions(+), 137 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 011fa34428c15..19befbc09fbe9 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -822,6 +822,13 @@ theorem HasFPowerSeriesOnBall.tendsto_partialSum Tendsto (fun n => p.partialSum n y) atTop (𝓝 (f (x + y))) := (hf.hasSum hy).tendsto_sum_nat +theorem HasFPowerSeriesAt.tendsto_partialSum + (hf : HasFPowerSeriesAt f p x) : + ∀ᶠ y in 𝓝 0, Tendsto (fun n => p.partialSum n y) atTop (𝓝 (f (x + y))) := by + rcases hf with ⟨r, hr⟩ + filter_upwards [EMetric.ball_mem_nhds (0 : E) hr.r_pos] with y hy + exact hr.tendsto_partialSum hy + open Finset in /-- If a function admits a power series expansion within a ball, then the partial sums `p.partialSum n z` converge to `f (x + y)` as `n → ∞` and `z → y`. Note that `x + z` doesn't need diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 08f88f4cedb1f..d9bd84e10d143 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -331,29 +331,34 @@ section variable (𝕜 E) /-- The identity formal multilinear series, with all coefficients equal to `0` except for `n = 1` -where it is (the continuous multilinear version of) the identity. -/ -def id : FormalMultilinearSeries 𝕜 E E - | 0 => 0 +where it is (the continuous multilinear version of) the identity. We allow an arbitrary +constant coefficient `x`. -/ +def id (x : E) : FormalMultilinearSeries 𝕜 E E + | 0 => ContinuousMultilinearMap.uncurry0 𝕜 _ x | 1 => (continuousMultilinearCurryFin1 𝕜 E E).symm (ContinuousLinearMap.id 𝕜 E) | _ => 0 +@[simp] theorem id_apply_zero (x : E) (v : Fin 0 → E) : + (FormalMultilinearSeries.id 𝕜 E x) 0 v = x := rfl + /-- The first coefficient of `id 𝕜 E` is the identity. -/ @[simp] -theorem id_apply_one (v : Fin 1 → E) : (FormalMultilinearSeries.id 𝕜 E) 1 v = v 0 := +theorem id_apply_one (x : E) (v : Fin 1 → E) : (FormalMultilinearSeries.id 𝕜 E x) 1 v = v 0 := rfl /-- The `n`th coefficient of `id 𝕜 E` is the identity when `n = 1`. We state this in a dependent way, as it will often appear in this form. -/ -theorem id_apply_one' {n : ℕ} (h : n = 1) (v : Fin n → E) : - (id 𝕜 E) n v = v ⟨0, h.symm ▸ zero_lt_one⟩ := by +theorem id_apply_one' (x : E) {n : ℕ} (h : n = 1) (v : Fin n → E) : + (id 𝕜 E x) n v = v ⟨0, h.symm ▸ zero_lt_one⟩ := by subst n apply id_apply_one /-- For `n ≠ 1`, the `n`-th coefficient of `id 𝕜 E` is zero, by definition. -/ @[simp] -theorem id_apply_ne_one {n : ℕ} (h : n ≠ 1) : (FormalMultilinearSeries.id 𝕜 E) n = 0 := by +theorem id_apply_of_one_lt (x : E) {n : ℕ} (h : 1 < n) : + (FormalMultilinearSeries.id 𝕜 E x) n = 0 := by cases' n with n - · rfl + · contradiction · cases n · contradiction · rfl @@ -361,11 +366,11 @@ theorem id_apply_ne_one {n : ℕ} (h : n ≠ 1) : (FormalMultilinearSeries.id end @[simp] -theorem comp_id (p : FormalMultilinearSeries 𝕜 E F) : p.comp (id 𝕜 E) = p := by +theorem comp_id (p : FormalMultilinearSeries 𝕜 E F) (x : E) : p.comp (id 𝕜 E x) = p := by ext1 n dsimp [FormalMultilinearSeries.comp] rw [Finset.sum_eq_single (Composition.ones n)] - · show compAlongComposition p (id 𝕜 E) (Composition.ones n) = p n + · show compAlongComposition p (id 𝕜 E x) (Composition.ones n) = p n ext v rw [compAlongComposition_apply] apply p.congr (Composition.ones_length n) @@ -375,50 +380,60 @@ theorem comp_id (p : FormalMultilinearSeries 𝕜 E F) : p.comp (id 𝕜 E) = p rw [Fin.ext_iff, Fin.coe_castLE, Fin.val_mk] · show ∀ b : Composition n, - b ∈ Finset.univ → b ≠ Composition.ones n → compAlongComposition p (id 𝕜 E) b = 0 + b ∈ Finset.univ → b ≠ Composition.ones n → compAlongComposition p (id 𝕜 E x) b = 0 intro b _ hb obtain ⟨k, hk, lt_k⟩ : ∃ (k : ℕ), k ∈ Composition.blocks b ∧ 1 < k := Composition.ne_ones_iff.1 hb obtain ⟨i, hi⟩ : ∃ (i : Fin b.blocks.length), b.blocks[i] = k := List.get_of_mem hk - let j : Fin b.length := ⟨i.val, b.blocks_length ▸ i.prop⟩ have A : 1 < b.blocksFun j := by convert lt_k ext v rw [compAlongComposition_apply, ContinuousMultilinearMap.zero_apply] apply ContinuousMultilinearMap.map_coord_zero _ j dsimp [applyComposition] - rw [id_apply_ne_one _ _ (ne_of_gt A)] + rw [id_apply_of_one_lt _ _ _ A] rfl · simp @[simp] -theorem id_comp (p : FormalMultilinearSeries 𝕜 E F) (h : p 0 = 0) : (id 𝕜 F).comp p = p := by +theorem id_comp (p : FormalMultilinearSeries 𝕜 E F) (v0 : Fin 0 → E) : + (id 𝕜 F (p 0 v0)).comp p = p := by ext1 n by_cases hn : n = 0 - · rw [hn, h] + · rw [hn] ext v - rw [comp_coeff_zero', id_apply_ne_one _ _ zero_ne_one] - rfl + simp only [comp_coeff_zero', id_apply_zero] + congr with i + exact i.elim0 · dsimp [FormalMultilinearSeries.comp] have n_pos : 0 < n := bot_lt_iff_ne_bot.mpr hn rw [Finset.sum_eq_single (Composition.single n n_pos)] - · show compAlongComposition (id 𝕜 F) p (Composition.single n n_pos) = p n + · show compAlongComposition (id 𝕜 F (p 0 v0)) p (Composition.single n n_pos) = p n ext v - rw [compAlongComposition_apply, id_apply_one' _ _ (Composition.single_length n_pos)] + rw [compAlongComposition_apply, id_apply_one' _ _ _ (Composition.single_length n_pos)] dsimp [applyComposition] refine p.congr rfl fun i him hin => congr_arg v <| ?_ ext; simp · show - ∀ b : Composition n, - b ∈ Finset.univ → b ≠ Composition.single n n_pos → compAlongComposition (id 𝕜 F) p b = 0 + ∀ b : Composition n, b ∈ Finset.univ → b ≠ Composition.single n n_pos → + compAlongComposition (id 𝕜 F (p 0 v0)) p b = 0 intro b _ hb - have A : b.length ≠ 1 := by simpa [Composition.eq_single_iff_length] using hb + have A : 1 < b.length := by + have : b.length ≠ 1 := by simpa [Composition.eq_single_iff_length] using hb + have : 0 < b.length := Composition.length_pos_of_pos b n_pos + omega ext v - rw [compAlongComposition_apply, id_apply_ne_one _ _ A] + rw [compAlongComposition_apply, id_apply_of_one_lt _ _ _ A] rfl · simp +/-- Variant of `id_comp` in which the zero coefficient is given by an equality hypothesis instead +of a definitional equality. Useful for rewriting or simplifying out in some situations. -/ +theorem id_comp' (p : FormalMultilinearSeries 𝕜 E F) (x : F) (v0 : Fin 0 → E) (h : x = p 0 v0) : + (id 𝕜 F x).comp p = p := by + simp [h] + /-! ### Summability properties of the composition of formal power series -/ @@ -634,11 +649,12 @@ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) /-- The auxiliary set corresponding to the composition of partial sums asymptotically contains all possible compositions. -/ -theorem compPartialSumTarget_tendsto_atTop : - Tendsto (fun N => compPartialSumTarget 0 N N) atTop atTop := by +theorem compPartialSumTarget_tendsto_prod_atTop : + Tendsto (fun (p : ℕ × ℕ) => compPartialSumTarget 0 p.1 p.2) atTop atTop := by apply Monotone.tendsto_atTop_finset · intro m n hmn a ha - have : ∀ i, i < m → i < n := fun i hi => lt_of_lt_of_le hi hmn + have : ∀ i, i < m.1 → i < n.1 := fun i hi => lt_of_lt_of_le hi hmn.1 + have : ∀ i, i < m.2 → i < n.2 := fun i hi => lt_of_lt_of_le hi hmn.2 aesop · rintro ⟨n, c⟩ simp only [mem_compPartialSumTarget_iff] @@ -650,29 +666,35 @@ theorem compPartialSumTarget_tendsto_atTop : apply hn simp only [Finset.mem_image_of_mem, Finset.mem_coe, Finset.mem_univ] +/-- The auxiliary set corresponding to the composition of partial sums asymptotically contains +all possible compositions. -/ +theorem compPartialSumTarget_tendsto_atTop : + Tendsto (fun N => compPartialSumTarget 0 N N) atTop atTop := by + apply Tendsto.comp compPartialSumTarget_tendsto_prod_atTop tendsto_atTop_diagonal + /-- Composing the partial sums of two multilinear series coincides with the sum over all compositions in `compPartialSumTarget 0 N N`. This is precisely the motivation for the definition of `compPartialSumTarget`. -/ theorem comp_partialSum (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) - (N : ℕ) (z : E) : - q.partialSum N (∑ i ∈ Finset.Ico 1 N, p i fun _j => z) = - ∑ i ∈ compPartialSumTarget 0 N N, q.compAlongComposition p i.2 fun _j => z := by + (M N : ℕ) (z : E) : + q.partialSum M (∑ i ∈ Finset.Ico 1 N, p i fun _j => z) = + ∑ i ∈ compPartialSumTarget 0 M N, q.compAlongComposition p i.2 fun _j => z := by -- we expand the composition, using the multilinearity of `q` to expand along each coordinate. suffices H : - (∑ n ∈ Finset.range N, + (∑ n ∈ Finset.range M, ∑ r ∈ Fintype.piFinset fun i : Fin n => Finset.Ico 1 N, q n fun i : Fin n => p (r i) fun _j => z) = - ∑ i ∈ compPartialSumTarget 0 N N, q.compAlongComposition p i.2 fun _j => z by + ∑ i ∈ compPartialSumTarget 0 M N, q.compAlongComposition p i.2 fun _j => z by simpa only [FormalMultilinearSeries.partialSum, ContinuousMultilinearMap.map_sum_finset] using H -- rewrite the first sum as a big sum over a sigma type, in the finset -- `compPartialSumTarget 0 N N` rw [Finset.range_eq_Ico, Finset.sum_sigma'] -- use `compChangeOfVariables_sum`, saying that this change of variables respects sums - apply compChangeOfVariables_sum 0 N N + apply compChangeOfVariables_sum 0 M N rintro ⟨k, blocks_fun⟩ H - apply congr _ (compChangeOfVariables_length 0 N N H).symm + apply congr _ (compChangeOfVariables_length 0 M N H).symm intros - rw [← compChangeOfVariables_blocksFun 0 N N H] + rw [← compChangeOfVariables_blocksFun 0 M N H] rfl end FormalMultilinearSeries diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index ccd47d8263a54..eadc577c91124 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -4,36 +4,43 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Analytic.Composition +import Mathlib.Analysis.Analytic.Linear /-! # Inverse of analytic functions We construct the left and right inverse of a formal multilinear series with invertible linear term, -we prove that they coincide and study their properties (notably convergence). +we prove that they coincide and study their properties (notably convergence). We deduce that the +inverse of an analytic partial homeomorphism is analytic. ## Main statements -* `p.leftInv i`: the formal left inverse of the formal multilinear series `p`, - for `i : E ≃L[𝕜] F` which coincides with `p₁`. -* `p.rightInv i`: the formal right inverse of the formal multilinear series `p`, - for `i : E ≃L[𝕜] F` which coincides with `p₁`. -* `p.leftInv_comp` says that `p.leftInv i` is indeed a left inverse to `p` when `p₁ = i`. -* `p.rightInv_comp` says that `p.rightInv i` is indeed a right inverse to `p` when `p₁ = i`. +* `p.leftInv i x`: the formal left inverse of the formal multilinear series `p`, with constant + coefficient `x`, for `i : E ≃L[𝕜] F` which coincides with `p₁`. +* `p.rightInv i x`: the formal right inverse of the formal multilinear series `p`, with constant + coefficient `x`, for `i : E ≃L[𝕜] F` which coincides with `p₁`. +* `p.leftInv_comp` says that `p.leftInv i x` is indeed a left inverse to `p` when `p₁ = i`. +* `p.rightInv_comp` says that `p.rightInv i x` is indeed a right inverse to `p` when `p₁ = i`. * `p.leftInv_eq_rightInv`: the two inverses coincide. * `p.radius_rightInv_pos_of_radius_pos`: if a power series has a positive radius of convergence, then so does its inverse. +* `PartialHomeomorph.hasFPowerSeriesAt_symm` shows that, if a partial homeomorph has a power series + `p` at a point, with invertible linear part, then the inverse also has a power series at the + image point, given by `p.leftInv`. -/ -open scoped Topology +open scoped Topology ENNReal open Finset Filter -namespace FormalMultilinearSeries +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] +namespace FormalMultilinearSeries /-! ### The left inverse of a formal multilinear series -/ @@ -51,27 +58,27 @@ term compensates the rest of the sum, using `i⁻¹` as an inverse to `p₁`. These formulas only make sense when the constant term `p₀` vanishes. The definition we give is general, but it ignores the value of `p₀`. -/ -noncomputable def leftInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : +noncomputable def leftInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : FormalMultilinearSeries 𝕜 F E - | 0 => 0 + | 0 => ContinuousMultilinearMap.uncurry0 𝕜 _ x | 1 => (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm | n + 2 => -∑ c : { c : Composition (n + 2) // c.length < n + 2 }, - (leftInv p i (c : Composition (n + 2)).length).compAlongComposition + (leftInv p i x (c : Composition (n + 2)).length).compAlongComposition (p.compContinuousLinearMap i.symm) c @[simp] -theorem leftInv_coeff_zero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.leftInv i 0 = 0 := by rw [leftInv] +theorem leftInv_coeff_zero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.leftInv i x 0 = ContinuousMultilinearMap.uncurry0 𝕜 _ x := by rw [leftInv] @[simp] -theorem leftInv_coeff_one (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.leftInv i 1 = (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm := by rw [leftInv] +theorem leftInv_coeff_one (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.leftInv i x 1 = (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm := by rw [leftInv] /-- The left inverse does not depend on the zeroth coefficient of a formal multilinear series. -/ -theorem leftInv_removeZero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.removeZero.leftInv i = p.leftInv i := by +theorem leftInv_removeZero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.removeZero.leftInv i x = p.leftInv i x := by ext1 n induction' n using Nat.strongRec' with n IH match n with @@ -87,14 +94,15 @@ theorem leftInv_removeZero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[ /-- The left inverse to a formal multilinear series is indeed a left inverse, provided its linear term is invertible. -/ -theorem leftInv_comp (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) - (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : (leftInv p i).comp p = id 𝕜 E := by - ext (n v) +theorem leftInv_comp (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) + (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : + (leftInv p i x).comp p = id 𝕜 E x := by + ext n v classical match n with | 0 => - simp only [leftInv_coeff_zero, ContinuousMultilinearMap.zero_apply, id_apply_ne_one, Ne, - not_false_iff, zero_ne_one, comp_coeff_zero'] + simp only [comp_coeff_zero', leftInv_coeff_zero, ContinuousMultilinearMap.uncurry0_apply, + id_apply_zero] | 1 => simp only [leftInv_coeff_one, comp_coeff_one, h, id_apply_one, ContinuousLinearEquiv.coe_apply, ContinuousLinearEquiv.symm_apply_apply, continuousMultilinearCurryFin1_symm_apply] @@ -111,16 +119,16 @@ theorem leftInv_comp (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) {Composition.ones (n + 2)} := by simp [Set.mem_toFinset (s := {c | Composition.length c < n + 2})] have C : - ((p.leftInv i (Composition.ones (n + 2)).length) + ((p.leftInv i x (Composition.ones (n + 2)).length) fun j : Fin (Composition.ones n.succ.succ).length => p 1 fun _ => v ((Fin.castLE (Composition.length_le _)) j)) = - p.leftInv i (n + 2) fun j : Fin (n + 2) => p 1 fun _ => v j := by + p.leftInv i x (n + 2) fun j : Fin (n + 2) => p 1 fun _ => v j := by apply FormalMultilinearSeries.congr _ (Composition.ones_length _) fun j hj1 hj2 => ?_ exact FormalMultilinearSeries.congr _ rfl fun k _ _ => by congr have D : - (p.leftInv i (n + 2) fun j : Fin (n + 2) => p 1 fun _ => v j) = + (p.leftInv i x (n + 2) fun j : Fin (n + 2) => p 1 fun _ => v j) = -∑ c ∈ {c : Composition (n + 2) | c.length < n + 2}.toFinset, - (p.leftInv i c.length) (p.applyComposition c v) := by + (p.leftInv i x c.length) (p.applyComposition c v) := by simp only [leftInv, ContinuousMultilinearMap.neg_apply, neg_inj, ContinuousMultilinearMap.sum_apply] convert @@ -128,7 +136,7 @@ theorem leftInv_comp (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (fun c : Composition (n + 2) => c.length < n + 2) (fun c : Composition (n + 2) => (ContinuousMultilinearMap.compAlongComposition - (p.compContinuousLinearMap (i.symm : F →L[𝕜] E)) c (p.leftInv i c.length)) + (p.compContinuousLinearMap (i.symm : F →L[𝕜] E)) c (p.leftInv i x c.length)) fun j : Fin (n + 2) => p 1 fun _ : Fin 1 => v j)).symm.trans _ simp only [compContinuousLinearMap_applyComposition, @@ -157,26 +165,26 @@ term compensates the rest of the sum, using `i⁻¹` as an inverse to `p₁`. These formulas only make sense when the constant term `p₀` vanishes. The definition we give is general, but it ignores the value of `p₀`. -/ -noncomputable def rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : +noncomputable def rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : FormalMultilinearSeries 𝕜 F E - | 0 => 0 + | 0 => ContinuousMultilinearMap.uncurry0 𝕜 _ x | 1 => (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm | n + 2 => - let q : FormalMultilinearSeries 𝕜 F E := fun k => if k < n + 2 then rightInv p i k else 0; + let q : FormalMultilinearSeries 𝕜 F E := fun k => if k < n + 2 then rightInv p i x k else 0; -(i.symm : F →L[𝕜] E).compContinuousMultilinearMap ((p.comp q) (n + 2)) @[simp] -theorem rightInv_coeff_zero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.rightInv i 0 = 0 := by rw [rightInv] +theorem rightInv_coeff_zero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.rightInv i x 0 = ContinuousMultilinearMap.uncurry0 𝕜 _ x := by rw [rightInv] @[simp] -theorem rightInv_coeff_one (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.rightInv i 1 = (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm := by rw [rightInv] +theorem rightInv_coeff_one (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.rightInv i x 1 = (continuousMultilinearCurryFin1 𝕜 F E).symm i.symm := by rw [rightInv] /-- The right inverse does not depend on the zeroth coefficient of a formal multilinear series. -/ -theorem rightInv_removeZero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) : - p.removeZero.rightInv i = p.rightInv i := by +theorem rightInv_removeZero (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) : + p.removeZero.rightInv i x = p.rightInv i x := by ext1 n induction' n using Nat.strongRec' with n IH match n with @@ -216,12 +224,12 @@ theorem comp_rightInv_aux1 {n : ℕ} (hn : 0 < n) (p : FormalMultilinearSeries simp [FormalMultilinearSeries.comp, A, Finset.sum_union B, C, -Set.toFinset_setOf, -add_right_inj, -Composition.single_length] -theorem comp_rightInv_aux2 (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (n : ℕ) +theorem comp_rightInv_aux2 (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) (n : ℕ) (v : Fin (n + 2) → F) : ∑ c ∈ {c : Composition (n + 2) | 1 < c.length}.toFinset, - p c.length (applyComposition (fun k : ℕ => ite (k < n + 2) (p.rightInv i k) 0) c v) = + p c.length (applyComposition (fun k : ℕ => ite (k < n + 2) (p.rightInv i x k) 0) c v) = ∑ c ∈ {c : Composition (n + 2) | 1 < c.length}.toFinset, - p c.length ((p.rightInv i).applyComposition c v) := by + p c.length ((p.rightInv i x).applyComposition c v) := by have N : 0 < n + 2 := by norm_num refine sum_congr rfl fun c hc => p.congr rfl fun j hj1 hj2 => ?_ have : ∀ k, c.blocksFun k < n + 2 := by @@ -232,14 +240,16 @@ theorem comp_rightInv_aux2 (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[ /-- The right inverse to a formal multilinear series is indeed a right inverse, provided its linear term is invertible and its constant term vanishes. -/ -theorem comp_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) - (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) (h0 : p 0 = 0) : - p.comp (rightInv p i) = id 𝕜 F := by +theorem comp_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) + (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : + p.comp (rightInv p i x) = id 𝕜 F (p 0 0) := by ext (n v) match n with | 0 => - simp only [h0, ContinuousMultilinearMap.zero_apply, id_apply_ne_one, Ne, not_false_iff, - zero_ne_one, comp_coeff_zero'] + simp only [comp_coeff_zero', Matrix.zero_empty, id_apply_zero] + congr + ext i + exact i.elim0 | 1 => simp only [comp_coeff_one, h, rightInv_coeff_one, ContinuousLinearEquiv.apply_symm_apply, id_apply_one, ContinuousLinearEquiv.coe_apply, continuousMultilinearCurryFin1_symm_apply] @@ -248,11 +258,12 @@ theorem comp_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F simp [comp_rightInv_aux1 N, h, rightInv, lt_irrefl n, show n + 2 ≠ 1 by omega, ← sub_eq_add_neg, sub_eq_zero, comp_rightInv_aux2, -Set.toFinset_setOf] -theorem rightInv_coeff (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (n : ℕ) (hn : 2 ≤ n) : - p.rightInv i n = +theorem rightInv_coeff (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) + (n : ℕ) (hn : 2 ≤ n) : + p.rightInv i x n = -(i.symm : F →L[𝕜] E).compContinuousMultilinearMap (∑ c ∈ ({c | 1 < Composition.length c}.toFinset : Finset (Composition n)), - p.compAlongComposition (p.rightInv i) c) := by + p.compAlongComposition (p.rightInv i x) c) := by match n with | 0 => exact False.elim (zero_lt_two.not_le hn) | 1 => exact False.elim (one_lt_two.not_le hn) @@ -267,26 +278,15 @@ theorem rightInv_coeff (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] /-! ### Coincidence of the left and the right inverse -/ -private theorem leftInv_eq_rightInv_aux (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) - (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) (h0 : p 0 = 0) : - leftInv p i = rightInv p i := - calc - leftInv p i = (leftInv p i).comp (id 𝕜 F) := by simp - _ = (leftInv p i).comp (p.comp (rightInv p i)) := by rw [comp_rightInv p i h h0] - _ = ((leftInv p i).comp p).comp (rightInv p i) := by rw [comp_assoc] - _ = (id 𝕜 E).comp (rightInv p i) := by rw [leftInv_comp p i h] - _ = rightInv p i := by simp - -/-- The left inverse and the right inverse of a formal multilinear series coincide. This is not at -all obvious from their definition, but it follows from uniqueness of inverses (which comes from the -fact that composition is associative on formal multilinear series). -/ -theorem leftInv_eq_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) - (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : leftInv p i = rightInv p i := +theorem leftInv_eq_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (x : E) + (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : + leftInv p i x = rightInv p i x := calc - leftInv p i = leftInv p.removeZero i := by rw [leftInv_removeZero] - _ = rightInv p.removeZero i := by - apply leftInv_eq_rightInv_aux _ _ (by simpa using h) (by simp) - _ = rightInv p i := by rw [rightInv_removeZero] + leftInv p i x = (leftInv p i x).comp (id 𝕜 F (p 0 0)) := by simp + _ = (leftInv p i x).comp (p.comp (rightInv p i x)) := by rw [comp_rightInv p i _ h] + _ = ((leftInv p i x).comp p).comp (rightInv p i x) := by rw [comp_assoc] + _ = (id 𝕜 E x).comp (rightInv p i x) := by rw [leftInv_comp p i _ h] + _ = rightInv p i x := by simp [id_comp' _ _ 0] /-! ### Convergence of the inverse of a power series @@ -423,17 +423,17 @@ theorem radius_right_inv_pos_of_radius_pos_aux1 (n : ℕ) (p : ℕ → ℝ) (hp expression for `∑_{k ‖p.rightInv i k‖) + radius_right_inv_pos_of_radius_pos_aux1 n (fun k => ‖p.rightInv i x k‖) (fun k => norm_nonneg _) hr ha /-- If a a formal multilinear series has a positive radius of convergence, then its right inverse also has a positive radius of convergence. -/ -theorem radius_rightInv_pos_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) - (hp : 0 < p.radius) : 0 < (p.rightInv i).radius := by +theorem radius_rightInv_pos_of_radius_pos + {p : FormalMultilinearSeries 𝕜 E F} {i : E ≃L[𝕜] F} {x : E} + (hp : 0 < p.radius) : 0 < (p.rightInv i x).radius := by obtain ⟨C, r, Cpos, rpos, ple⟩ : ∃ (C r : _) (_ : 0 < C) (_ : 0 < r), ∀ n : ℕ, ‖p n‖ ≤ C * r ^ n := le_mul_pow_of_radius_pos p hp @@ -508,7 +505,7 @@ theorem radius_rightInv_pos_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) exact ⟨a, ha.1, ha.2.1.le, ha.2.2.le⟩ -- check by induction that the partial sums are suitably bounded, using the choice of `a` and the -- inductive control from Lemma `radius_rightInv_pos_of_radius_pos_aux2`. - let S n := ∑ k ∈ Ico 1 n, a ^ k * ‖p.rightInv i k‖ + let S n := ∑ k ∈ Ico 1 n, a ^ k * ‖p.rightInv i x k‖ have IRec : ∀ n, 1 ≤ n → S n ≤ (I + 1) * a := by apply Nat.le_induction · simp only [S] @@ -536,21 +533,159 @@ theorem radius_rightInv_pos_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) _ ≤ (I + 1) * a := by gcongr -- conclude that all coefficients satisfy `aⁿ Qₙ ≤ (I + 1) a`. let a' : NNReal := ⟨a, apos.le⟩ - suffices H : (a' : ENNReal) ≤ (p.rightInv i).radius by + suffices H : (a' : ENNReal) ≤ (p.rightInv i x).radius by apply lt_of_lt_of_le _ H -- Prior to leanprover/lean4#2734, this was `exact_mod_cast apos`. simpa only [ENNReal.coe_pos] - apply le_radius_of_bound _ ((I + 1) * a) fun n => ?_ - by_cases hn : n = 0 - · have : ‖p.rightInv i n‖ = ‖p.rightInv i 0‖ := by congr <;> try rw [hn] - simp only [this, norm_zero, zero_mul, rightInv_coeff_zero] - positivity - · have one_le_n : 1 ≤ n := bot_lt_iff_ne_bot.2 hn - calc - ‖p.rightInv i n‖ * (a' : ℝ) ^ n = a ^ n * ‖p.rightInv i n‖ := mul_comm _ _ - _ ≤ ∑ k ∈ Ico 1 (n + 1), a ^ k * ‖p.rightInv i k‖ := - (haveI : ∀ k ∈ Ico 1 (n + 1), 0 ≤ a ^ k * ‖p.rightInv i k‖ := fun k _ => by positivity - single_le_sum this (by simp [one_le_n])) - _ ≤ (I + 1) * a := IRec (n + 1) (by norm_num) + apply le_radius_of_eventually_le _ ((I + 1) * a) + filter_upwards [Ici_mem_atTop 1] with n (hn : 1 ≤ n) + calc + ‖p.rightInv i x n‖ * (a' : ℝ) ^ n = a ^ n * ‖p.rightInv i x n‖ := mul_comm _ _ + _ ≤ ∑ k ∈ Ico 1 (n + 1), a ^ k * ‖p.rightInv i x k‖ := + (haveI : ∀ k ∈ Ico 1 (n + 1), 0 ≤ a ^ k * ‖p.rightInv i x k‖ := fun k _ => by positivity + single_le_sum this (by simp [hn])) + _ ≤ (I + 1) * a := IRec (n + 1) (by norm_num) + +/-- If a a formal multilinear series has a positive radius of convergence, then its left inverse +also has a positive radius of convergence. -/ +theorem radius_leftInv_pos_of_radius_pos + {p : FormalMultilinearSeries 𝕜 E F} {i : E ≃L[𝕜] F} {x : E} + (hp : 0 < p.radius) (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : + 0 < (p.leftInv i x).radius := by + rw [leftInv_eq_rightInv _ _ _ h] + exact radius_rightInv_pos_of_radius_pos hp end FormalMultilinearSeries + +/-! +### The inverse of an analytic partial homeomorphism is analytic +-/ + +open FormalMultilinearSeries List + +lemma HasFPowerSeriesAt.tendsto_partialSum_prod_of_comp + {f : E → G} {q : FormalMultilinearSeries 𝕜 F G} + {p : FormalMultilinearSeries 𝕜 E F} {x : E} + (hf : HasFPowerSeriesAt f (q.comp p) x) (hq : 0 < q.radius) (hp : 0 < p.radius) : + ∀ᶠ y in 𝓝 0, Tendsto (fun (a : ℕ × ℕ) ↦ q.partialSum a.1 (p.partialSum a.2 y + - p 0 (fun _ ↦ 0))) atTop (𝓝 (f (x + y))) := by + rcases hf with ⟨r0, h0⟩ + rcases q.comp_summable_nnreal p hq hp with ⟨r1, r1_pos : 0 < r1, hr1⟩ + let r : ℝ≥0∞ := min r0 r1 + have : EMetric.ball (0 : E) r ∈ 𝓝 0 := + EMetric.ball_mem_nhds 0 (lt_min h0.r_pos (by exact_mod_cast r1_pos)) + filter_upwards [this] with y hy + have hy0 : y ∈ EMetric.ball 0 r0 := EMetric.ball_subset_ball (min_le_left _ _) hy + have A : HasSum (fun i : Σ n, Composition n => q.compAlongComposition p i.2 fun _j => y) + (f (x + y)) := by + have cau : CauchySeq fun s : Finset (Σ n, Composition n) => + ∑ i ∈ s, q.compAlongComposition p i.2 fun _j => y := by + apply cauchySeq_finset_of_norm_bounded _ (NNReal.summable_coe.2 hr1) _ + simp only [coe_nnnorm, NNReal.coe_mul, NNReal.coe_pow] + rintro ⟨n, c⟩ + calc + ‖(compAlongComposition q p c) fun _j : Fin n => y‖ ≤ + ‖compAlongComposition q p c‖ * ∏ _j : Fin n, ‖y‖ := by + apply ContinuousMultilinearMap.le_opNorm + _ ≤ ‖compAlongComposition q p c‖ * (r1 : ℝ) ^ n := by + apply mul_le_mul_of_nonneg_left _ (norm_nonneg _) + rw [Finset.prod_const, Finset.card_fin] + apply pow_le_pow_left (norm_nonneg _) + rw [EMetric.mem_ball, edist_eq_coe_nnnorm] at hy + have := le_trans (le_of_lt hy) (min_le_right _ _) + rwa [ENNReal.coe_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this + apply HasSum.of_sigma (fun b ↦ hasSum_fintype _) ?_ cau + simpa [FormalMultilinearSeries.comp] using h0.hasSum hy0 + have B : Tendsto (fun (n : ℕ × ℕ) => ∑ i ∈ compPartialSumTarget 0 n.1 n.2, + q.compAlongComposition p i.2 fun _j => y) atTop (𝓝 (f (x + y))) := by + apply Tendsto.comp A compPartialSumTarget_tendsto_prod_atTop + have C : Tendsto (fun (n : ℕ × ℕ) => q.partialSum n.1 (∑ a ∈ Finset.Ico 1 n.2, p a fun _b ↦ y)) + atTop (𝓝 (f (x + y))) := by simpa [comp_partialSum] using B + apply C.congr' + filter_upwards [Ici_mem_atTop (0, 1)] + rintro ⟨-, n⟩ ⟨-, (hn : 1 ≤ n)⟩ + congr + rw [partialSum, eq_sub_iff_add_eq', Finset.range_eq_Ico, + Finset.sum_eq_sum_Ico_succ_bot hn] + congr with i + exact i.elim0 + +lemma HasFPowerSeriesAt.eventually_hasSum_of_comp {f : E → F} {g : F → G} + {q : FormalMultilinearSeries 𝕜 F G} {p : FormalMultilinearSeries 𝕜 E F} {x : E} + (hgf : HasFPowerSeriesAt (g ∘ f) (q.comp p) x) (hf : HasFPowerSeriesAt f p x) + (hq : 0 < q.radius) : + ∀ᶠ y in 𝓝 0, HasSum (fun n : ℕ => q n fun _ : Fin n => (f (x + y) - f x)) (g (f (x + y))) := by + have : ∀ᶠ y in 𝓝 (0 : E), f (x + y) - f x ∈ EMetric.ball 0 q.radius := by + have A : ContinuousAt (fun y ↦ f (x + y) - f x) 0 := by + apply ContinuousAt.sub _ continuousAt_const + exact hf.continuousAt.comp_of_eq (continuous_add_left x).continuousAt (by simp) + have B : EMetric.ball 0 q.radius ∈ 𝓝 (f (x + 0) - f x) := by + simpa using EMetric.ball_mem_nhds _ hq + exact A.preimage_mem_nhds B + filter_upwards [hgf.tendsto_partialSum_prod_of_comp hq (hf.radius_pos), + hf.tendsto_partialSum, this] with y hy h'y h''y + have L : Tendsto (fun n ↦ q.partialSum n (f (x + y) - f x)) atTop (𝓝 (g (f (x + y)))) := by + apply (closed_nhds_basis (g (f (x + y)))).tendsto_right_iff.2 + rintro u ⟨hu, u_closed⟩ + 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 + 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 + (𝓝 (q.partialSum a (f (x + y) - f x))) := by + have : ContinuousAt (q.partialSum a) (f (x + y) - f x) := + (partialSum_continuous q a).continuousAt + apply this.tendsto.comp + apply Tendsto.sub h'y + convert tendsto_const_nhds + exact (HasFPowerSeriesAt.coeff_zero hf fun _ ↦ 0).symm + apply u_closed.mem_of_tendsto this + filter_upwards [Ici_mem_atTop b₀] with b hb using vu (hab _ _ ha hb) + have C : CauchySeq (fun (s : Finset ℕ) ↦ ∑ n ∈ s, q n fun _ : Fin n => (f (x + y) - f x)) := by + have Z := q.summable_norm_apply (x := f (x + y) - f x) h''y + exact cauchySeq_finset_of_norm_bounded _ Z (fun i ↦ le_rfl) + exact tendsto_nhds_of_cauchySeq_of_subseq C tendsto_finset_range L + +/-- If a partial homeomorphism `f` is defined at `a` and has a power series expansion there with +invertible linear term, then `f.symm` has a power series expansion at `f a`, given by the inverse +of the initial power series. -/ +theorem PartialHomeomorph.hasFPowerSeriesAt_symm (f : PartialHomeomorph E F) {a : E} + {i : E ≃L[𝕜] F} (h0 : a ∈ f.source) {p : FormalMultilinearSeries 𝕜 E F} + (h : HasFPowerSeriesAt f p a) (hp : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : + HasFPowerSeriesAt f.symm (p.leftInv i a) (f a) := by + have A : HasFPowerSeriesAt (f.symm ∘ f) ((p.leftInv i a).comp p) a := by + have : HasFPowerSeriesAt (ContinuousLinearMap.id 𝕜 E) ((p.leftInv i a).comp p) a := by + rw [leftInv_comp _ _ _ hp] + exact (ContinuousLinearMap.id 𝕜 E).hasFPowerSeriesAt a + apply this.congr + filter_upwards [f.open_source.mem_nhds h0] with x hx using by simp [hx] + have B : ∀ᶠ (y : E) in 𝓝 0, HasSum (fun n ↦ (p.leftInv i a n) fun _ ↦ f (a + y) - f a) + (f.symm (f (a + y))) := by + simpa using A.eventually_hasSum_of_comp h (radius_leftInv_pos_of_radius_pos h.radius_pos hp) + have C : ∀ᶠ (y : E) in 𝓝 a, HasSum (fun n ↦ (p.leftInv i a n) fun _ ↦ f y - f a) + (f.symm (f y)) := by + rw [← sub_eq_zero_of_eq (a := a) rfl] at B + have : ContinuousAt (fun x ↦ x - a) a := by fun_prop + simpa using this.preimage_mem_nhds B + have D : ∀ᶠ (y : E) in 𝓝 (f.symm (f a)), + HasSum (fun n ↦ (p.leftInv i a n) fun _ ↦ f y - f a) y := by + simp only [h0, PartialHomeomorph.left_inv] + filter_upwards [C, f.open_source.mem_nhds h0] with x hx h'x + simpa [h'x] using hx + have E : ∀ᶠ z in 𝓝 (f a), HasSum (fun n ↦ (p.leftInv i a n) fun _ ↦ f (f.symm z) - f a) + (f.symm z) := by + have : ContinuousAt f.symm (f a) := f.continuousAt_symm (f.map_source h0) + exact this D + have F : ∀ᶠ z in 𝓝 (f a), HasSum (fun n ↦ (p.leftInv i a n) fun _ ↦ z - f a) (f.symm z) := by + filter_upwards [f.open_target.mem_nhds (f.map_source h0), E] with z hz h'z + simpa [hz] using h'z + rcases EMetric.mem_nhds_iff.1 F with ⟨r, r_pos, hr⟩ + refine ⟨min r (p.leftInv i a).radius, min_le_right _ _, + lt_min r_pos (radius_leftInv_pos_of_radius_pos h.radius_pos hp), fun {y} hy ↦ ?_⟩ + have : y + f a ∈ EMetric.ball (f a) r := by + simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, sub_zero, lt_min_iff, + add_sub_cancel_right] at hy ⊢ + exact hy.1 + simpa [add_comm] using hr this diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index c68372a5f2ce0..b47e95b56f4ad 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -146,7 +146,35 @@ end ContinuousMul section CompleteSpace -variable [CommGroup α] [UniformSpace α] [UniformGroup α] [CompleteSpace α] +variable [CommGroup α] [UniformSpace α] [UniformGroup α] + +@[to_additive] +theorem HasProd.of_sigma {γ : β → Type*} {f : (Σ b : β, γ b) → α} {g : β → α} {a : α} + (hf : ∀ b, HasProd (fun c ↦ f ⟨b, c⟩) (g b)) (hg : HasProd g a) + (h : CauchySeq (fun (s : Finset (Σ b : β, γ b)) ↦ ∏ i ∈ s, f i)) : + HasProd f a := by + classical + apply le_nhds_of_cauchy_adhp h + simp only [← mapClusterPt_def, mapClusterPt_iff, frequently_atTop, ge_iff_le, le_eq_subset] + intro u hu s + rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, hv⟩ + 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 + (𝓝 <| ∏ 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 := + ((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⟩ + simpa only [mem_filter, st hx, true_and] using ht0 (mem_image_of_mem Sigma.fst hx) + +variable [CompleteSpace α] @[to_additive] theorem Multipliable.sigma_factor {γ : β → Type*} {f : (Σb : β, γ b) → α} From 488d6b54f90af375e6297be294a365284f0a756d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 28 Sep 2024 09:50:06 +0000 Subject: [PATCH 005/174] feat(Topology): define `IsOpenQuotientMap` (#16549) --- Mathlib.lean | 1 + Mathlib/Topology/Constructions.lean | 4 ++ Mathlib/Topology/Defs/Basic.lean | 22 +++++++++ Mathlib/Topology/Maps/OpenQuotient.lean | 60 +++++++++++++++++++++++++ 4 files changed, 87 insertions(+) create mode 100644 Mathlib/Topology/Maps/OpenQuotient.lean diff --git a/Mathlib.lean b/Mathlib.lean index 161a4867d386f..891f2a03ada8d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4654,6 +4654,7 @@ import Mathlib.Topology.LocallyConstant.Algebra import Mathlib.Topology.LocallyConstant.Basic import Mathlib.Topology.LocallyFinite import Mathlib.Topology.Maps.Basic +import Mathlib.Topology.Maps.OpenQuotient import Mathlib.Topology.Maps.Proper.Basic import Mathlib.Topology.Maps.Proper.UniversallyClosed import Mathlib.Topology.MetricSpace.Algebra diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index d8f4038358755..f23b56ce022b3 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -817,6 +817,10 @@ theorem embedding_graph {f : X → Y} (hf : Continuous f) : Embedding fun x => ( theorem embedding_prod_mk (x : X) : Embedding (Prod.mk x : Y → X × Y) := embedding_of_embedding_compose (Continuous.Prod.mk x) continuous_snd embedding_id +theorem IsOpenQuotientMap.prodMap {f : X → Y} {g : Z → W} (hf : IsOpenQuotientMap f) + (hg : IsOpenQuotientMap g) : IsOpenQuotientMap (Prod.map f g) := + ⟨.prodMap hf.1 hg.1, .prod_map hf.2 hg.2, .prod hf.3 hg.3⟩ + end Prod section Bool diff --git a/Mathlib/Topology/Defs/Basic.lean b/Mathlib/Topology/Defs/Basic.lean index 607362753bace..e7526dd6e9024 100644 --- a/Mathlib/Topology/Defs/Basic.lean +++ b/Mathlib/Topology/Defs/Basic.lean @@ -151,6 +151,28 @@ def IsOpenMap (f : X → Y) : Prop := ∀ U : Set X, IsOpen U → IsOpen (f '' U if the image of any closed `U : Set X` is closed in `Y`. -/ def IsClosedMap (f : X → Y) : Prop := ∀ U : Set X, IsClosed U → IsClosed (f '' U) +/-- An open quotient map is an open map `f : X → Y` which is both an open map and a quotient map. +Equivalently, it is a surjective continuous open map. +We use the latter characterization as a definition. + +Many important quotient maps are open quotient maps, including + +- the quotient map from a topological space to its quotient by the action of a group; +- the quotient map from a topological group to its quotient by a normal subgroup; +- the quotient map from a topological spaace to its separation quotient. + +Contrary to general quotient maps, +the category of open quotient maps is closed under `Prod.map`. +-/ +@[mk_iff] +structure IsOpenQuotientMap (f : X → Y) : Prop where + /-- An open quotient map is surjective. -/ + surjective : Function.Surjective f + /-- An open quotient map is continuous. -/ + continuous : Continuous f + /-- An open quotient map is an open map. -/ + isOpenMap : IsOpenMap f + end Defs /-! ### Notation for non-standard topologies -/ diff --git a/Mathlib/Topology/Maps/OpenQuotient.lean b/Mathlib/Topology/Maps/OpenQuotient.lean new file mode 100644 index 0000000000000..f5dc76959e06b --- /dev/null +++ b/Mathlib/Topology/Maps/OpenQuotient.lean @@ -0,0 +1,60 @@ +/- +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.Maps.Basic + +/-! +# Open quotient maps + +An open quotient map is an open map `f : X → Y` which is both an open map and a quotient map. +Equivalently, it is a surjective continuous open map. +We use the latter characterization as a definition. + +Many important quotient maps are open quotient maps, including + +- the quotient map from a topological space to its quotient by the action of a group; +- the quotient map from a topological group to its quotient by a normal subgroup; +- the quotient map from a topological spaace to its separation quotient. + +Contrary to general quotient maps, +the category of open quotient maps is closed under `Prod.map`. +-/ + +open Function Filter +open scoped Topology + +variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : X → Y} + +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 iff_isOpenMap_quotientMap : IsOpenQuotientMap f ↔ IsOpenMap f ∧ QuotientMap f := + ⟨fun h ↦ ⟨h.isOpenMap, h.quotientMap⟩, fun ⟨ho, hq⟩ ↦ ⟨hq.surjective, hq.continuous, ho⟩⟩ + +theorem of_isOpenMap_quotientMap (ho : IsOpenMap f) (hq : QuotientMap f) : + IsOpenQuotientMap f := + iff_isOpenMap_quotientMap.2 ⟨ho, hq⟩ + +theorem comp {g : Y → Z} (hg : IsOpenQuotientMap g) (hf : IsOpenQuotientMap f) : + IsOpenQuotientMap (g ∘ f) := + ⟨.comp hg.1 hf.1, .comp hg.2 hf.2, .comp hg.3 hf.3⟩ + +theorem map_nhds_eq (h : IsOpenQuotientMap f) (x : X) : map f (𝓝 x) = 𝓝 (f x) := + le_antisymm h.continuous.continuousAt <| h.isOpenMap.nhds_le _ + +theorem continuous_comp_iff (h : IsOpenQuotientMap f) {g : Y → Z} : + Continuous (g ∘ f) ↔ Continuous g := + h.quotientMap.continuous_iff.symm + +theorem continuousAt_comp_iff (h : IsOpenQuotientMap f) {g : Y → Z} {x : X} : + ContinuousAt (g ∘ f) x ↔ ContinuousAt g (f x) := by + simp only [ContinuousAt, ← h.map_nhds_eq, tendsto_map'_iff, comp_def] + +end IsOpenQuotientMap From 68c55d08790260870b5981b403f5921bbbbb9a0b Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sat, 28 Sep 2024 10:21:34 +0000 Subject: [PATCH 006/174] chore(FieldTheory/Galois): move file Galois into folder Galois (#17197) Move file Galois into folder Galois, rename as basic for the incomming infinite case. Moves: - Mathlib.FieldTheory.Galois -> Mathlib.FieldTheory.Galois.Basic --- Mathlib.lean | 2 +- Mathlib/FieldTheory/Finite/GaloisField.lean | 2 +- Mathlib/FieldTheory/{Galois.lean => Galois/Basic.lean} | 0 Mathlib/FieldTheory/IsSepClosed.lean | 2 +- Mathlib/FieldTheory/KrullTopology.lean | 2 +- Mathlib/FieldTheory/KummerExtension.lean | 2 +- Mathlib/FieldTheory/PolynomialGaloisGroup.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Basic.lean | 2 +- Mathlib/RingTheory/Norm/Basic.lean | 2 +- Mathlib/RingTheory/Trace/Basic.lean | 2 +- 10 files changed, 9 insertions(+), 9 deletions(-) rename Mathlib/FieldTheory/{Galois.lean => Galois/Basic.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 891f2a03ada8d..256900385c5ac 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2660,7 +2660,7 @@ import Mathlib.FieldTheory.Finite.Polynomial import Mathlib.FieldTheory.Finite.Trace import Mathlib.FieldTheory.Finiteness import Mathlib.FieldTheory.Fixed -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic import Mathlib.FieldTheory.IntermediateField.Algebraic import Mathlib.FieldTheory.IntermediateField.Basic import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure diff --git a/Mathlib/FieldTheory/Finite/GaloisField.lean b/Mathlib/FieldTheory/Finite/GaloisField.lean index baca752d4d5c6..d26ad82dcaec4 100644 --- a/Mathlib/FieldTheory/Finite/GaloisField.lean +++ b/Mathlib/FieldTheory/Finite/GaloisField.lean @@ -6,7 +6,7 @@ Authors: Aaron Anderson, Alex J. Best, Johan Commelin, Eric Rodriguez, Ruben Van import Mathlib.Algebra.CharP.Algebra import Mathlib.Data.ZMod.Algebra import Mathlib.FieldTheory.Finite.Basic -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic import Mathlib.FieldTheory.SplittingField.IsSplittingField /-! diff --git a/Mathlib/FieldTheory/Galois.lean b/Mathlib/FieldTheory/Galois/Basic.lean similarity index 100% rename from Mathlib/FieldTheory/Galois.lean rename to Mathlib/FieldTheory/Galois/Basic.lean diff --git a/Mathlib/FieldTheory/IsSepClosed.lean b/Mathlib/FieldTheory/IsSepClosed.lean index a986c4144c374..1a11e8d38dff7 100644 --- a/Mathlib/FieldTheory/IsSepClosed.lean +++ b/Mathlib/FieldTheory/IsSepClosed.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic /-! # Separably Closed Field diff --git a/Mathlib/FieldTheory/KrullTopology.lean b/Mathlib/FieldTheory/KrullTopology.lean index 18d9b02736673..7f7f53c3109a7 100644 --- a/Mathlib/FieldTheory/KrullTopology.lean +++ b/Mathlib/FieldTheory/KrullTopology.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Sebastian Monnet. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Monnet -/ -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic import Mathlib.Topology.Algebra.FilterBasis import Mathlib.Topology.Algebra.OpenSubgroup import Mathlib.Tactic.ByContra diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 18680d5a83dd6..4be37a02bc718 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -5,7 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.RingTheory.RootsOfUnity.Basic import Mathlib.RingTheory.AdjoinRoot -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic import Mathlib.LinearAlgebra.Eigenspace.Minpoly import Mathlib.RingTheory.Norm.Basic /-! diff --git a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean index 159fca2636398..003d87b0e80d0 100644 --- a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean +++ b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Thomas Browning, Patrick Lutz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz -/ -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic /-! # Galois Groups of Polynomials diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 891a8f3be52b6..20d24352136cb 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -5,7 +5,7 @@ Authors: Riccardo Brasca -/ import Mathlib.RingTheory.Polynomial.Cyclotomic.Roots import Mathlib.NumberTheory.NumberField.Basic -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic /-! # Cyclotomic extensions diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index a051d47766d54..98235e2bf157f 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -8,7 +8,7 @@ import Mathlib.FieldTheory.PrimitiveElement import Mathlib.LinearAlgebra.Matrix.Charpoly.Minpoly import Mathlib.LinearAlgebra.Matrix.ToLinearEquiv import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic /-! # Norm for (finite) ring extensions diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index ee8a5a44c08f7..6690ed277ff91 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.RingTheory.Trace.Defs import Mathlib.LinearAlgebra.Determinant -import Mathlib.FieldTheory.Galois +import Mathlib.FieldTheory.Galois.Basic import Mathlib.LinearAlgebra.Matrix.Charpoly.Minpoly import Mathlib.LinearAlgebra.Vandermonde import Mathlib.FieldTheory.Minpoly.MinpolyDiv From 06eff6ae7f695371103c2115593ba1e1cd4d494b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 28 Sep 2024 11:27:20 +0000 Subject: [PATCH 007/174] chore(Multilinear): golf (#17198) Also add 2 missing `simp` lemmas --- Mathlib/Data/Finset/Pi.lean | 4 ++++ Mathlib/Data/Fintype/Pi.lean | 9 +++++---- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 11 +++-------- 3 files changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/Finset/Pi.lean b/Mathlib/Data/Finset/Pi.lean index 8edf7868f430a..9751986c4519c 100644 --- a/Mathlib/Data/Finset/Pi.lean +++ b/Mathlib/Data/Finset/Pi.lean @@ -92,6 +92,10 @@ lemma pi_nonempty : (s.pi t).Nonempty ↔ ∀ a ∈ s, (t a).Nonempty := by @[aesop safe apply (rule_sets := [finsetNonempty])] alias ⟨_, pi_nonempty_of_forall_nonempty⟩ := pi_nonempty +@[simp] +lemma pi_eq_empty : s.pi t = ∅ ↔ ∃ a ∈ s, t a = ∅ := by + simp [← not_nonempty_iff_eq_empty] + @[simp] theorem pi_insert [∀ a, DecidableEq (β a)] {s : Finset α} {t : ∀ a : α, Finset (β a)} {a : α} (ha : a ∉ s) : pi (insert a s) t = (t a).biUnion fun b => (pi s t).image (Pi.cons s a b) := by diff --git a/Mathlib/Data/Fintype/Pi.lean b/Mathlib/Data/Fintype/Pi.lean index 8b57b9e376500..99a88bda7004a 100644 --- a/Mathlib/Data/Fintype/Pi.lean +++ b/Mathlib/Data/Fintype/Pi.lean @@ -49,12 +49,13 @@ theorem piFinset_subset (t₁ t₂ : ∀ a, Finset (δ a)) (h : ∀ a, t₁ a piFinset t₁ ⊆ piFinset t₂ := fun _ hg => mem_piFinset.2 fun a => h a <| mem_piFinset.1 hg a @[simp] -theorem piFinset_empty [Nonempty α] : piFinset (fun _ => ∅ : ∀ i, Finset (δ i)) = ∅ := - eq_empty_of_forall_not_mem fun _ => by simp +theorem piFinset_eq_empty : piFinset s = ∅ ↔ ∃ i, s i = ∅ := by simp [piFinset] @[simp] -lemma piFinset_nonempty : (piFinset s).Nonempty ↔ ∀ a, (s a).Nonempty := by - simp [Finset.Nonempty, Classical.skolem] +theorem piFinset_empty [Nonempty α] : piFinset (fun _ => ∅ : ∀ i, Finset (δ i)) = ∅ := by simp + +@[simp] +lemma piFinset_nonempty : (piFinset s).Nonempty ↔ ∀ a, (s a).Nonempty := by simp [piFinset] @[aesop safe apply (rule_sets := [finsetNonempty])] alias ⟨_, Aesop.piFinset_nonempty_of_forall_nonempty⟩ := piFinset_nonempty diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index dbc2bea4ba67f..2a78e4ae591af 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -453,14 +453,9 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, induction' n using Nat.strong_induction_on with n IH generalizing A -- If one of the sets is empty, then all the sums are zero by_cases Ai_empty : ∃ i, A i = ∅ - · rcases Ai_empty with ⟨i, hi⟩ - have : ∑ j ∈ A i, g i j = 0 := by rw [hi, Finset.sum_empty] - rw [f.map_coord_zero i this] - have : piFinset A = ∅ := by - refine Finset.eq_empty_of_forall_not_mem fun r hr => ?_ - have : r i ∈ A i := mem_piFinset.mp hr i - simp [hi] at this - rw [this, Finset.sum_empty] + · obtain ⟨i, hi⟩ : ∃ i, ∑ j ∈ A i, g i j = 0 := Ai_empty.imp fun i hi ↦ by simp [hi] + have hpi : piFinset A = ∅ := by simpa + rw [f.map_coord_zero i hi, hpi, Finset.sum_empty] push_neg at Ai_empty -- Otherwise, if all sets are at most singletons, then they are exactly singletons and the result -- is again straightforward From 6b353d7c2caa00c2d9d97b9a6c644fc668f19fb0 Mon Sep 17 00:00:00 2001 From: 97l Date: Sat, 28 Sep 2024 14:33:18 +0000 Subject: [PATCH 008/174] feat: Binomial theorem for subtraction (#17128) Co-authored-by: ahhwuhu Co-authored-by: 97l <130048468+ahhwuhu@users.noreply.github.com> --- Mathlib/Data/Nat/Choose/Sum.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index d9bafecbbe840..f339da88c6daf 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -74,6 +74,17 @@ theorem add_pow [CommSemiring R] (x y : R) (n : ℕ) : (x + y) ^ n = ∑ m ∈ range (n + 1), x ^ m * y ^ (n - m) * n.choose m := (Commute.all x y).add_pow n +/-- A special case of the **binomial theorem** -/ +theorem sub_pow [CommRing R] (x y : R) (n : ℕ) : + (x - y) ^ n = ∑ m ∈ range (n + 1), (-1) ^ (m + n) * x ^ m * y ^ (n - m) * n.choose m := by + rw [sub_eq_add_neg, add_pow] + congr! 1 with m hm + have : (-1 : R) ^ (n - m) = (-1) ^ (n + m) := by + rw [mem_range] at hm + simp [show n + m = n - m + 2 * m by omega, pow_add] + rw [neg_pow, this] + ring + namespace Nat /-- The sum of entries in a row of Pascal's triangle -/ From 9d7e2cd94764b3dd4b04932930b308bd531c6f19 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 15:08:13 +0000 Subject: [PATCH 009/174] =?UTF-8?q?feat:=20`dens=20s=E2=81=BB=C2=B9=20=3D?= =?UTF-8?q?=20dens=20s`=20(#17202)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanAPAP --- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index b36ab5a81b89b..a44299063bf7c 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -272,6 +272,9 @@ theorem coe_inv (s : Finset α) : ↑s⁻¹ = (s : Set α)⁻¹ := coe_image.tra @[to_additive (attr := simp)] theorem card_inv (s : Finset α) : s⁻¹.card = s.card := card_image_of_injective _ inv_injective +@[to_additive (attr := simp)] +lemma dens_inv [Fintype α] (s : Finset α) : s⁻¹.dens = s.dens := by simp [dens] + @[to_additive (attr := simp)] theorem preimage_inv (s : Finset α) : s.preimage (·⁻¹) inv_injective.injOn = s⁻¹ := coe_injective <| by rw [coe_preimage, Set.inv_preimage, coe_inv] From ebef379baba79e98d8dca59983ca172dcc4c2dc2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 15:08:14 +0000 Subject: [PATCH 010/174] feat: coercion from a submonoid is a Freiman homomorphism (#17209) From LeanAPAP --- Mathlib/Combinatorics/Additive/FreimanHom.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index c196e507f58bb..6b1ff91dbf629 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Bhavik Mehta import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Order.BigOperators.Group.Multiset import Mathlib.Data.ZMod.Defs @@ -249,6 +250,11 @@ lemma isMulFreimanIso_empty : IsMulFreimanIso n (∅ : Set α) (∅ : Set β) f map_prod_eq_map_prod s t _ _ _ _ := by rw [← map_multiset_prod, ← map_multiset_prod, EquivLike.apply_eq_iff_eq] +@[to_additive] +lemma IsMulFreimanHom.subtypeVal {S : Type*} [SetLike S α] [SubmonoidClass S α] {s : S} : + IsMulFreimanHom n (univ : Set s) univ Subtype.val := + MonoidHomClass.isMulFreimanHom (SubmonoidClass.subtype s) (mapsTo_univ ..) + end CommMonoid section CancelCommMonoid From fada5f49bb517a689c349119dd0a19fb762d95af Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 15:08:15 +0000 Subject: [PATCH 011/174] feat(Data/Finset): cardinality of a preimage (#17211) From LeanAPAP --- Mathlib/Data/Finset/Preimage.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 840e7e2c93e3e..742274880b947 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -80,6 +80,10 @@ theorem map_subset_iff_subset_preimage {f : α ↪ β} {s : Finset α} {t : Fins s.map f ⊆ t ↔ s ⊆ t.preimage f f.injective.injOn := by classical rw [map_eq_image, image_subset_iff_subset_preimage] +lemma card_preimage (s : Finset β) (f : α → β) (hf) [DecidablePred (· ∈ Set.range f)] : + (s.preimage f hf).card = {x ∈ s | x ∈ Set.range f}.card := + card_nbij f (by simp) (by simpa) (fun b hb ↦ by aesop) + theorem image_preimage [DecidableEq β] (f : α → β) (s : Finset β) [∀ x, Decidable (x ∈ Set.range f)] (hf : Set.InjOn f (f ⁻¹' ↑s)) : image f (preimage s f hf) = s.filter fun x => x ∈ Set.range f := Finset.coe_inj.1 <| by From 96fdd12ab1170884ba721729d5009d1a61bba2b7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 28 Sep 2024 16:01:37 +0000 Subject: [PATCH 012/174] feat(Alternating/Topology): add `ClosedEmbedding` (#17200) Add `ContinuousAlternatingMap.isClosed_range_toContinuousMultilinearMap` and `ContinuousAlternatingMap.closedEmbedding_toContinuousMultilinearMap` --- .../Algebra/Module/Alternating/Topology.lean | 31 ++++++++++++++----- 1 file changed, 23 insertions(+), 8 deletions(-) diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 988d2a453721f..747f1739055fe 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -24,6 +24,22 @@ namespace ContinuousAlternatingMap variable {𝕜 E F ι : Type*} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [AddCommGroup F] [Module 𝕜 F] +section IsClosedRange + +variable [TopologicalSpace F] [TopologicalAddGroup F] + +instance instTopologicalSpace : TopologicalSpace (E [⋀^ι]→L[𝕜] F) := + .induced toContinuousMultilinearMap inferInstance + +lemma isClosed_range_toContinuousMultilinearMap [ContinuousSMul 𝕜 E] [T2Space F] : + IsClosed (Set.range (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → + ContinuousMultilinearMap 𝕜 (fun _ : ι ↦ E) F)) := by + simp only [range_toContinuousMultilinearMap, setOf_forall] + repeat refine isClosed_iInter fun _ ↦ ?_ + exact isClosed_singleton.preimage (ContinuousMultilinearMap.continuous_eval_const _) + +end IsClosedRange + section UniformAddGroup variable [UniformSpace F] [UniformAddGroup F] @@ -67,11 +83,8 @@ theorem completeSpace (h : RestrictGenTopology {s : Set (ι → E) | IsVonNBound CompleteSpace (E [⋀^ι]→L[𝕜] F) := by have := ContinuousMultilinearMap.completeSpace (F := F) h rw [completeSpace_iff_isComplete_range - uniformEmbedding_toContinuousMultilinearMap.toUniformInducing, range_toContinuousMultilinearMap] - simp only [setOf_forall] - apply IsClosed.isComplete - repeat refine isClosed_iInter fun _ ↦ ?_ - exact isClosed_singleton.preimage (ContinuousMultilinearMap.continuous_eval_const _) + uniformEmbedding_toContinuousMultilinearMap.toUniformInducing] + apply isClosed_range_toContinuousMultilinearMap.isComplete instance instCompleteSpace [TopologicalAddGroup E] [SequentialSpace (ι → E)] : CompleteSpace (E [⋀^ι]→L[𝕜] F) := @@ -100,9 +113,6 @@ end UniformAddGroup variable [TopologicalSpace F] [TopologicalAddGroup F] -instance instTopologicalSpace : TopologicalSpace (E [⋀^ι]→L[𝕜] F) := - .induced toContinuousMultilinearMap inferInstance - lemma embedding_toContinuousMultilinearMap : Embedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F → _)) := letI := TopologicalAddGroup.toUniformSpace F @@ -138,6 +148,11 @@ theorem hasBasis_nhds_zero : variable [ContinuousSMul 𝕜 E] +lemma closedEmbedding_toContinuousMultilinearMap [T2Space F] : + ClosedEmbedding (toContinuousMultilinearMap : + (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 := From 01252d1a6f839f6087e70a512e98d1f050a9d0a1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 16:01:38 +0000 Subject: [PATCH 013/174] =?UTF-8?q?feat:=20`x=20^=20(log=20x)=E2=81=BB?= =?UTF-8?q?=C2=B9=20=E2=89=A4=20exp=201`=20(#17203)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and `-x⁻¹ ≤ log x ≤ x`. All three statements are weakenings of more natural statements that hold only when `x ≠ 0`. Having versions that also work for `x = 0` due to junk values is really practical. From LeanAPAP --- Mathlib/Analysis/SpecialFunctions/Log/Basic.lean | 10 ++++++++++ Mathlib/Analysis/SpecialFunctions/Pow/Real.lean | 16 ++++++++++++++++ 2 files changed, 26 insertions(+) diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 4698f6d30f954..705e2df157be4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -281,6 +281,16 @@ theorem log_le_sub_one_of_pos {x : ℝ} (hx : 0 < x) : log x ≤ x - 1 := by lemma one_sub_inv_le_log_of_pos (hx : 0 < x) : 1 - x⁻¹ ≤ log x := by simpa [add_comm] using log_le_sub_one_of_pos (inv_pos.2 hx) +/-- See `Real.log_le_sub_one_of_pos` for the stronger version when `x ≠ 0`. -/ +lemma log_le_self (hx : 0 ≤ x) : log x ≤ x := by + obtain rfl | hx := hx.eq_or_lt + · simp + · exact (log_le_sub_one_of_pos hx).trans (by linarith) + +/-- See `Real.one_sub_inv_le_log_of_pos` for the stronger version when `x ≠ 0`. -/ +lemma neg_inv_le_log (hx : 0 ≤ x) : -x⁻¹ ≤ log x := by + rw [neg_le, ← log_inv]; exact log_le_self <| inv_nonneg.2 hx + /-- Bound for `|log x * x|` in the interval `(0, 1]`. -/ theorem abs_log_mul_self_lt (x : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) : |log x * x| < 1 := by have : 0 < 1 / x := by simpa only [one_div, inv_pos] using h1 diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 6a7ea43f88beb..0bfe01f772dee 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -166,6 +166,22 @@ theorem abs_rpow_le_exp_log_mul (x y : ℝ) : |x ^ y| ≤ exp (log x * y) := by · by_cases hy : y = 0 <;> simp [hx, hy, zero_le_one] · rw [rpow_def_of_pos (abs_pos.2 hx), log_abs] +lemma rpow_inv_log (hx₀ : 0 < x) (hx₁ : x ≠ 1) : x ^ (log x)⁻¹ = exp 1 := by + rw [rpow_def_of_pos hx₀, mul_inv_cancel₀] + exact log_ne_zero.2 ⟨hx₀.ne', hx₁, (hx₀.trans' <| by norm_num).ne'⟩ + +/-- See `Real.rpow_inv_log` for the equality when `x ≠ 1` is strictly positive. -/ +lemma rpow_inv_log_le_exp_one : x ^ (log x)⁻¹ ≤ exp 1 := by + calc + _ ≤ |x ^ (log x)⁻¹| := le_abs_self _ + _ ≤ |x| ^ (log x)⁻¹ := abs_rpow_le_abs_rpow .. + rw [← log_abs] + obtain hx | hx := (abs_nonneg x).eq_or_gt + · simp [hx] + · rw [rpow_def_of_pos hx] + gcongr + exact mul_inv_le_one + theorem norm_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : ‖x ^ y‖ = ‖x‖ ^ y := by simp_rw [Real.norm_eq_abs] exact abs_rpow_of_nonneg hx_nonneg From 3fdf47759894d7a29c0b5891a6bc55ddf468f95c Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Sat, 28 Sep 2024 16:22:42 +0000 Subject: [PATCH 014/174] feat(RingTheory/Ideal/KrullDimension): define the Krull dimension of a ring (#14134) The ring theoretic krull dimension is the order theoretic krull dimension applied to its prime spectrum. Unfolding this definition, it is the length of longest series of prime ideals ordered by inclusion. Co-authored-by: Matthew Ballard Co-authored-by: Andrew Yang Co-authored-by: mariainesdff Co-authored-by: FMLJohn <128468682+FMLJohn@users.noreply.github.com> --- Mathlib.lean | 2 + .../PrimeSpectrum/Basic.lean | 14 +++- Mathlib/Order/KrullDimension.lean | 10 +++ Mathlib/RingTheory/KrullDimension/Basic.lean | 64 +++++++++++++++++++ Mathlib/RingTheory/KrullDimension/Field.lean | 22 +++++++ Mathlib/Topology/KrullDimension.lean | 7 +- scripts/noshake.json | 12 ++-- 7 files changed, 117 insertions(+), 14 deletions(-) create mode 100644 Mathlib/RingTheory/KrullDimension/Basic.lean create mode 100644 Mathlib/RingTheory/KrullDimension/Field.lean diff --git a/Mathlib.lean b/Mathlib.lean index 256900385c5ac..94bbbdbd563f3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3923,6 +3923,8 @@ import Mathlib.RingTheory.JacobsonIdeal import Mathlib.RingTheory.Kaehler.Basic import Mathlib.RingTheory.Kaehler.CotangentComplex import Mathlib.RingTheory.Kaehler.Polynomial +import Mathlib.RingTheory.KrullDimension.Basic +import Mathlib.RingTheory.KrullDimension.Field import Mathlib.RingTheory.LaurentSeries import Mathlib.RingTheory.LittleWedderburn import Mathlib.RingTheory.LocalProperties.Basic diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 766ddec99f38a..c12db5f61af3a 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -3,9 +3,8 @@ 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.RingTheory.PrimeSpectrum -import Mathlib.Topology.Irreducible -import Mathlib.Topology.Sets.Closeds +import Mathlib.RingTheory.KrullDimension.Basic +import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober import Mathlib.RingTheory.Ideal.MinimalPrime import Mathlib.RingTheory.Ideal.Over @@ -655,3 +654,12 @@ theorem PrimeSpectrum.comap_residue (T : Type u) [CommRing T] [LocalRing T] exact Ideal.mk_ker end LocalRing + +section KrullDimension + +theorem PrimeSpectrum.topologicalKrullDim_eq_ringKrullDim [CommRing R] : + topologicalKrullDim (PrimeSpectrum R) = ringKrullDim R := + Order.krullDim_orderDual.symm.trans <| Order.krullDim_eq_of_orderIso + (PrimeSpectrum.pointsEquivIrreducibleCloseds R).symm + +end KrullDimension diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 86ed7e2ce6e89..4631e6250c4ba 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -174,6 +174,8 @@ variable {α β : Type*} variable [Preorder α] [Preorder β] +lemma LTSeries.length_le_krullDim (p : LTSeries α) : p.length ≤ krullDim α := le_sSup ⟨_, rfl⟩ + lemma krullDim_nonneg_of_nonempty [Nonempty α] : 0 ≤ krullDim α := le_sSup ⟨⟨0, fun _ ↦ @Nonempty.some α inferInstance, fun f ↦ f.elim0⟩, rfl⟩ @@ -214,6 +216,14 @@ lemma krullDim_eq_zero_of_unique [Unique α] : krullDim α = 0 := by by_contra r exact ne_of_lt (q.step ⟨0, not_le.mp r⟩) <| Subsingleton.elim _ _ +lemma krullDim_nonpos_of_subsingleton [Subsingleton α] : krullDim α ≤ 0 := by + by_cases hα : Nonempty α + · have := uniqueOfSubsingleton (Classical.choice hα) + exact le_of_eq krullDim_eq_zero_of_unique + · have := not_nonempty_iff.mp hα + exact le_of_lt <| lt_of_eq_of_lt krullDim_eq_bot_of_isEmpty <| + Batteries.compareOfLessAndEq_eq_lt.mp rfl + lemma krullDim_le_of_strictComono_and_surj (f : α → β) (hf : ∀ ⦃a b⦄, f a < f b → a < b) (hf' : Function.Surjective f) : krullDim β ≤ krullDim α := diff --git a/Mathlib/RingTheory/KrullDimension/Basic.lean b/Mathlib/RingTheory/KrullDimension/Basic.lean new file mode 100644 index 0000000000000..cca9f211a88e5 --- /dev/null +++ b/Mathlib/RingTheory/KrullDimension/Basic.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Fangming Li, Jujian Zhang +-/ +import Mathlib.Algebra.MvPolynomial.CommRing +import Mathlib.Algebra.Polynomial.Basic +import Mathlib.RingTheory.PrimeSpectrum +import Mathlib.Order.KrullDimension + +/-! +# Krull dimensions of (commutative) rings + +Given a commutative ring, its ring theoretic Krull dimension is the order theoretic Krull dimension +of its prime spectrum. Unfolding this definition, it is the length of the longest sequence(s) of +prime ideals ordered by strict inclusion. +-/ + +open Order + +/-- +The ring theoretic Krull dimension is the Krull dimension of its spectrum ordered by inclusion. +-/ +noncomputable def ringKrullDim (R : Type*) [CommRing R] : WithBot (WithTop ℕ) := + krullDim (PrimeSpectrum R) + +variable {R S : Type*} [CommRing R] [CommRing S] + +@[nontriviality] +lemma ringKrullDim_eq_bot_of_subsingleton [Subsingleton R] : + ringKrullDim R = ⊥ := + krullDim_eq_bot_of_isEmpty + +lemma ringKrullDim_nonneg_of_nontrivial [Nontrivial R] : + 0 ≤ ringKrullDim R := + krullDim_nonneg_of_nonempty + +/-- If `f : R →+* S` is surjective, then `ringKrullDim S ≤ ringKrullDim R`. -/ +theorem ringKrullDim_le_of_surjective (f : R →+* S) (hf : Function.Surjective f) : + ringKrullDim S ≤ ringKrullDim R := + krullDim_le_of_strictMono (fun I ↦ ⟨Ideal.comap f I.asIdeal, inferInstance⟩) + (Monotone.strictMono_of_injective (fun _ _ hab ↦ Ideal.comap_mono hab) + (fun _ _ h => PrimeSpectrum.ext_iff.mpr <| Ideal.comap_injective_of_surjective f hf <| by + simpa using h)) + +/-- If `I` is an ideal of `R`, then `ringKrullDim (R ⧸ I) ≤ ringKrullDim R`. -/ +theorem ringKrullDim_quotient_le (I : Ideal R) : + ringKrullDim (R ⧸ I) ≤ ringKrullDim R := + ringKrullDim_le_of_surjective _ Ideal.Quotient.mk_surjective + +/-- If `R` and `S` are isomorphic, then `ringKrullDim R = ringKrullDim S`. -/ +theorem ringKrullDim_eq_of_ringEquiv (e : R ≃+* S) : + ringKrullDim R = ringKrullDim S := + le_antisymm (ringKrullDim_le_of_surjective e.symm e.symm.surjective) + (ringKrullDim_le_of_surjective e e.surjective) + +alias RingEquiv.ringKrullDim := ringKrullDim_eq_of_ringEquiv + +proof_wanted Polynomial.ringKrullDim_le : + ringKrullDim (Polynomial R) ≤ 2 * (ringKrullDim R) + 1 + +proof_wanted MvPolynomial.fin_ringKrullDim_eq_add_of_isNoetherianRing + [IsNoetherianRing R] (n : ℕ) : + ringKrullDim (MvPolynomial (Fin n) R) = ringKrullDim R + n diff --git a/Mathlib/RingTheory/KrullDimension/Field.lean b/Mathlib/RingTheory/KrullDimension/Field.lean new file mode 100644 index 0000000000000..bc8ebdf09a387 --- /dev/null +++ b/Mathlib/RingTheory/KrullDimension/Field.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Fangming Li, Jujian Zhang +-/ +import Mathlib.RingTheory.KrullDimension.Basic + +/-! +# The Krull dimension of a field + +This file proves that the Krull dimension of a field is zero. +-/ + +open Order + +@[simp] +theorem ringKrullDim_eq_zero_of_field (F : Type*) [Field F] : ringKrullDim F = 0 := + krullDim_eq_zero_of_unique + +theorem ringKrullDim_eq_zero_of_isField {F : Type*} [CommRing F] (hF : IsField F) : + ringKrullDim F = 0 := + @krullDim_eq_zero_of_unique _ _ <| @PrimeSpectrum.instUnique _ hF.toField diff --git a/Mathlib/Topology/KrullDimension.lean b/Mathlib/Topology/KrullDimension.lean index c76906356f367..ea4e77fb0194c 100644 --- a/Mathlib/Topology/KrullDimension.lean +++ b/Mathlib/Topology/KrullDimension.lean @@ -12,16 +12,13 @@ import Mathlib.Topology.Sets.Closeds The Krull dimension of a topological space is the order theoretic Krull dimension applied to the collection of all its subsets that are closed and irreducible. Unfolding this definition, it is the length of longest series of closed irreducible subsets ordered by inclusion. - -TODO: The Krull dimension of `Spec(R)` equals the Krull dimension of `R`, for `R` a commutative - ring. -/ -open TopologicalSpace +open TopologicalSpace Order /-- The Krull dimension of a topological space is the supremum of lengths of chains of closed irreducible sets. -/ noncomputable def topologicalKrullDim (T : Type*) [TopologicalSpace T] : WithBot ℕ∞ := - Order.krullDim (IrreducibleCloseds T) + krullDim (IrreducibleCloseds T) diff --git a/scripts/noshake.json b/scripts/noshake.json index f7776819fcd27..5335397efa8fc 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -240,8 +240,7 @@ ["Mathlib.Tactic.LinearCombination.Lemmas"], "Mathlib.Tactic.Lemma": ["Lean.Parser.Command"], "Mathlib.Tactic.IrreducibleDef": ["Mathlib.Data.Subtype"], - "Mathlib.Tactic.ITauto": - ["Batteries.Tactic.Init", "Mathlib.Logic.Basic"], + "Mathlib.Tactic.ITauto": ["Batteries.Tactic.Init", "Mathlib.Logic.Basic"], "Mathlib.Tactic.Group": ["Mathlib.Algebra.Group.Commutator"], "Mathlib.Tactic.GCongr.Core": ["Mathlib.Order.Defs"], "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], @@ -307,6 +306,8 @@ "Mathlib.RingTheory.PolynomialAlgebra": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.RingTheory.MvPolynomial.Homogeneous": ["Mathlib.Algebra.DirectSum.Internal"], + "Mathlib.RingTheory.KrullDimension.Basic": + ["Mathlib.Algebra.MvPolynomial.CommRing", "Mathlib.Algebra.Polynomial.Basic"], "Mathlib.RingTheory.Binomial": ["Mathlib.Algebra.Order.Floor"], "Mathlib.RepresentationTheory.FdRep": ["Mathlib.CategoryTheory.Monoidal.Rigid.Braided"], @@ -364,7 +365,6 @@ "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Basic": ["Mathlib.Data.Finset.Attr"], - "Mathlib.Data.DFinsupp.Notation": ["Mathlib.Data.Finsupp.Notation"], "Mathlib.Data.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Bool.Basic": ["Batteries.Tactic.Init"], "Mathlib.Control.Traversable.Instances": ["Mathlib.Control.Applicative"], @@ -382,13 +382,13 @@ ["Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback", "Mathlib.CategoryTheory.Limits.Shapes.Pullbacks"], "Mathlib.CategoryTheory.Limits.IsLimit": ["Batteries.Tactic.Congr"], - "Mathlib.CategoryTheory.Bicategory.Adjunction": - ["Mathlib.Tactic.CategoryTheory.Bicategory.Basic"], + "Mathlib.CategoryTheory.Category.Basic": ["Mathlib.Tactic.StacksAttribute"], "Mathlib.CategoryTheory.Bicategory.Functor.Oplax": ["Mathlib.Tactic.CategoryTheory.ToApp"], "Mathlib.CategoryTheory.Bicategory.Functor.Lax": ["Mathlib.Tactic.CategoryTheory.ToApp"], - "Mathlib.CategoryTheory.Category.Basic": ["Mathlib.Tactic.StacksAttribute"], + "Mathlib.CategoryTheory.Bicategory.Adjunction": + ["Mathlib.Tactic.CategoryTheory.Bicategory.Basic"], "Mathlib.Analysis.Normed.Operator.LinearIsometry": ["Mathlib.Algebra.Star.Basic"], "Mathlib.Analysis.InnerProductSpace.Basic": From 67904bc4822c8f4dafda24391cfd01f27c23f721 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 16:22:43 +0000 Subject: [PATCH 015/174] chore: Make `Finset.univ_nonempty` be simp (#17216) `Set.univ_nonempty` already is From LeanAPAP --- Mathlib/Algebra/BigOperators/Expect.lean | 2 +- Mathlib/Data/Fintype/Basic.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Expect.lean b/Mathlib/Algebra/BigOperators/Expect.lean index ed859ddca94e4..b64a24d549e6c 100644 --- a/Mathlib/Algebra/BigOperators/Expect.lean +++ b/Mathlib/Algebra/BigOperators/Expect.lean @@ -416,7 +416,7 @@ end AddCommMonoid section Semiring variable [Semiring M] [Module ℚ≥0 M] -@[simp] lemma expect_one [Nonempty ι] : 𝔼 _i : ι, (1 : M) = 1 := expect_const _ +lemma expect_one [Nonempty ι] : 𝔼 _i : ι, (1 : M) = 1 := expect_const _ end Semiring end Fintype diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 1852e40e4fd39..0ff3b973f0141 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -93,7 +93,7 @@ theorem Nonempty.eq_univ [Subsingleton α] : s.Nonempty → s = univ := by theorem univ_nonempty_iff : (univ : Finset α).Nonempty ↔ Nonempty α := by rw [← coe_nonempty, coe_univ, Set.nonempty_iff_univ_nonempty] -@[aesop unsafe apply (rule_sets := [finsetNonempty])] +@[simp, aesop unsafe apply (rule_sets := [finsetNonempty])] theorem univ_nonempty [Nonempty α] : (univ : Finset α).Nonempty := univ_nonempty_iff.2 ‹_› From 5cd5bfff8a4b7d1ba91e8374b1e149e72b26b8ca Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 28 Sep 2024 17:13:43 +0000 Subject: [PATCH 016/174] chore: rename Submodule.coeSubtype (#17219) --- Mathlib/Algebra/DirectSum/Module.lean | 4 ++-- Mathlib/Algebra/Lie/EngelSubalgebra.lean | 2 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 2 +- Mathlib/Algebra/Lie/Weights/Cartan.lean | 4 ++-- Mathlib/Algebra/Module/Submodule/LinearMap.lean | 4 +++- Mathlib/Algebra/Module/Submodule/Pointwise.lean | 2 +- Mathlib/Algebra/Module/ZLattice/Basic.lean | 2 +- Mathlib/Algebra/Star/Module.lean | 4 ++-- Mathlib/Analysis/Calculus/Implicit.lean | 4 ++-- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 2 +- Mathlib/LinearAlgebra/Basis/Basic.lean | 2 +- Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean | 4 ++-- Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean | 2 +- Mathlib/LinearAlgebra/DFinsupp.lean | 4 ++-- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 2 +- Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean | 2 +- Mathlib/LinearAlgebra/Finsupp.lean | 2 +- Mathlib/LinearAlgebra/FreeModule/PID.lean | 4 ++-- Mathlib/LinearAlgebra/Isomorphisms.lean | 4 ++-- Mathlib/LinearAlgebra/Projection.lean | 4 ++-- Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean | 8 ++++---- Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean | 8 ++++---- Mathlib/NumberTheory/NumberField/FractionalIdeal.lean | 2 +- Mathlib/Probability/ConditionalExpectation.lean | 4 ++-- Mathlib/RingTheory/Flat/Basic.lean | 2 +- 25 files changed, 43 insertions(+), 41 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index bda16093f4ab9..291bd79759b1b 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -279,7 +279,7 @@ theorem coeLinearMap_eq_dfinsupp_sum [DecidableEq M] (x : DirectSum ι fun i => simp only [coeLinearMap, toModule, DFinsupp.lsum, LinearEquiv.coe_mk, LinearMap.coe_mk, AddHom.coe_mk] rw [DFinsupp.sumAddHom_apply] - simp only [LinearMap.toAddMonoidHom_coe, Submodule.coeSubtype] + simp only [LinearMap.toAddMonoidHom_coe, Submodule.coe_subtype] @[simp] theorem coeLinearMap_of (i : ι) (x : A i) : DirectSum.coeLinearMap A (of (fun i ↦ A i) i x) = x := @@ -349,7 +349,7 @@ theorem IsInternal.collectedBasis_coe (h : IsInternal A) {α : ι → Type*} simp only [DFinsupp.mapRange_single, Basis.repr_symm_apply, linearCombination_single, one_smul, toModule] erw [DFinsupp.lsum_single] - simp only [Submodule.coeSubtype] + simp only [Submodule.coe_subtype] theorem IsInternal.collectedBasis_mem (h : IsInternal A) {α : ι → Type*} (v : ∀ i, Basis (α i) R (A i)) (a : Σi, α i) : h.collectedBasis v a ∈ A a.1 := by simp diff --git a/Mathlib/Algebra/Lie/EngelSubalgebra.lean b/Mathlib/Algebra/Lie/EngelSubalgebra.lean index 7257a4abad970..6720922dae1f2 100644 --- a/Mathlib/Algebra/Lie/EngelSubalgebra.lean +++ b/Mathlib/Algebra/Lie/EngelSubalgebra.lean @@ -137,7 +137,7 @@ lemma normalizer_eq_self_of_engel_le [IsArtinian R L] rintro _ ⟨y, rfl⟩ simp only [pow_succ', LinearMap.mul_apply, Submodule.mem_comap, mem_coe_submodule] apply aux₁ - simp only [Submodule.coeSubtype, SetLike.coe_mem] + simp only [Submodule.coe_subtype, SetLike.coe_mem] /-- A Lie subalgebra of a Noetherian Lie algebra is nilpotent if it is contained in the Engel subalgebra of all its elements. -/ diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 405111118ee13..ada8114afb6da 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -141,7 +141,7 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where apply add_mem -- Now `⁅a, y⁆ ∈ J` since `a ∈ I`, `y ∈ J`, and `J` is an ideal of `I`. · simp only [Submodule.mem_map, LieSubmodule.mem_coeSubmodule, Subtype.exists] - erw [Submodule.coeSubtype] + erw [Submodule.coe_subtype] simp only [exists_and_right, exists_eq_right, ha, lie_mem_left, exists_true_left] exact lie_mem_right R I J ⟨a, ha⟩ y hy -- Finally `⁅b, y⁆ = 0`, by the independence of the atoms. diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index bdc8524325ed0..0ef812f460f9e 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -289,9 +289,9 @@ lemma mem_corootSpace' {x : H} : erw [← (H : Submodule R L).injective_subtype.mem_set_image (s := Submodule.span R s)] rw [mem_image] simp_rw [SetLike.mem_coe] - rw [← Submodule.mem_map, Submodule.coeSubtype, Submodule.map_span, mem_corootSpace, ← this] + rw [← Submodule.mem_map, Submodule.coe_subtype, Submodule.map_span, mem_corootSpace, ← this] ext u - simp only [Submodule.coeSubtype, mem_image, Subtype.exists, LieSubalgebra.mem_coe_submodule, + simp only [Submodule.coe_subtype, mem_image, Subtype.exists, LieSubalgebra.mem_coe_submodule, exists_and_right, exists_eq_right, mem_setOf_eq, s] refine ⟨fun ⟨_, y, hy, z, hz, hyz⟩ ↦ ⟨y, hy, z, hz, hyz⟩, fun ⟨y, hy, z, hz, hyz⟩ ↦ ⟨?_, y, hy, z, hz, hyz⟩⟩ diff --git a/Mathlib/Algebra/Module/Submodule/LinearMap.lean b/Mathlib/Algebra/Module/Submodule/LinearMap.lean index 4ff64ac81d3b1..913de1199bb0c 100644 --- a/Mathlib/Algebra/Module/Submodule/LinearMap.lean +++ b/Mathlib/Algebra/Module/Submodule/LinearMap.lean @@ -75,9 +75,11 @@ theorem subtype_apply (x : p) : p.subtype x = x := rfl @[simp] -theorem coeSubtype : (Submodule.subtype p : p → M) = Subtype.val := +theorem coe_subtype : (Submodule.subtype p : p → M) = Subtype.val := rfl +@[deprecated (since := "2024-09-27")] alias coeSubtype := coe_subtype + theorem injective_subtype : Injective p.subtype := Subtype.coe_injective diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index 031be85372fc7..c8d8a18665f39 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -418,7 +418,7 @@ lemma set_smul_eq_map [SMulCommClass R R N] : exact ⟨Finsupp.single r ⟨n, hn⟩, Finsupp.single_mem_supported _ _ hr, by simp⟩ · intro x hx obtain ⟨c, hc, rfl⟩ := hx - simp only [LinearMap.coe_comp, coeSubtype, Finsupp.coe_lsum, Finsupp.sum, Function.comp_apply] + simp only [LinearMap.coe_comp, coe_subtype, Finsupp.coe_lsum, Finsupp.sum, Function.comp_apply] rw [AddSubmonoid.coe_finset_sum] refine Submodule.sum_mem (p := sR • N) (t := c.support) ?_ _ ⟨sR • N, ?_⟩ · rintro r hr diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index 94a9e5badef5b..0ae1e6609864d 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -583,7 +583,7 @@ theorem Basis.ofZLatticeBasis_repr_apply (x : L) (i : ι) : exact DFunLike.congr_fun (LinearMap.congr_fun this x) i refine Basis.ext b fun i ↦ ?_ simp_rw [LinearMap.coe_comp, Function.comp_apply, LinearMap.coe_restrictScalars, - LinearEquiv.coe_coe, coeSubtype, ← b.ofZLatticeBasis_apply K, repr_self, + LinearEquiv.coe_coe, coe_subtype, ← b.ofZLatticeBasis_apply K, repr_self, Finsupp.mapRange.linearMap_apply, Finsupp.mapRange_single, Algebra.linearMap_apply, map_one] theorem Basis.ofZLatticeBasis_span : diff --git a/Mathlib/Algebra/Star/Module.lean b/Mathlib/Algebra/Star/Module.lean index b062492e6acbd..1640c1f29979a 100644 --- a/Mathlib/Algebra/Star/Module.lean +++ b/Mathlib/Algebra/Star/Module.lean @@ -184,8 +184,8 @@ def StarModule.decomposeProdAdjoint : A ≃ₗ[R] selfAdjoint A × skewAdjoint A refine LinearEquiv.ofLinear ((selfAdjointPart R).prod (skewAdjointPart R)) (LinearMap.coprod ((selfAdjoint.submodule R A).subtype) (skewAdjoint.submodule R A).subtype) ?_ (LinearMap.ext <| StarModule.selfAdjointPart_add_skewAdjointPart R) - -- Note: with #6965 `Submodule.coeSubtype` doesn't fire in `dsimp` or `simp` - ext x <;> dsimp <;> erw [Submodule.coeSubtype, Submodule.coeSubtype] <;> simp + -- Note: with #6965 `Submodule.coe_subtype` doesn't fire in `dsimp` or `simp` + ext x <;> dsimp <;> erw [Submodule.coe_subtype, Submodule.coe_subtype] <;> simp end SelfSkewAdjoint diff --git a/Mathlib/Analysis/Calculus/Implicit.lean b/Mathlib/Analysis/Calculus/Implicit.lean index 72620c2d28d46..2e323d8872a82 100644 --- a/Mathlib/Analysis/Calculus/Implicit.lean +++ b/Mathlib/Analysis/Calculus/Implicit.lean @@ -326,12 +326,12 @@ theorem to_implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : · ext -- Porting note: added parentheses to help `simp` simp only [Classical.choose_spec hker, implicitFunctionDataOfComplemented, - ContinuousLinearMap.comp_apply, Submodule.coe_subtypeL', Submodule.coeSubtype, + ContinuousLinearMap.comp_apply, Submodule.coe_subtypeL', Submodule.coe_subtype, ContinuousLinearMap.id_apply] swap · ext -- Porting note: added parentheses to help `simp` - simp only [(ContinuousLinearMap.comp_apply), Submodule.coe_subtypeL', Submodule.coeSubtype, + simp only [(ContinuousLinearMap.comp_apply), Submodule.coe_subtypeL', Submodule.coe_subtype, LinearMap.map_coe_ker, (ContinuousLinearMap.zero_apply)] simp only [implicitFunctionDataOfComplemented, map_sub, sub_self] diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 8bf622dbce6f5..03d1f51b135c5 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -947,7 +947,7 @@ theorem LinearIsometry.extend_apply (L : S →ₗᵢ[𝕜] V) (s : S) : L.extend simp only [add_right_eq_self, LinearIsometry.coe_toLinearMap, LinearIsometryEquiv.coe_toLinearIsometry, LinearIsometry.coe_comp, Function.comp_apply, orthogonalProjection_mem_subspace_eq_self, LinearMap.coe_comp, ContinuousLinearMap.coe_coe, - Submodule.coeSubtype, LinearMap.add_apply, Submodule.coe_eq_zero, + Submodule.coe_subtype, LinearMap.add_apply, Submodule.coe_eq_zero, LinearIsometryEquiv.map_eq_zero_iff, Submodule.coe_subtypeₗᵢ, orthogonalProjection_mem_subspace_orthogonalComplement_eq_zero, Submodule.orthogonal_orthogonal, Submodule.coe_mem] diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index 67204421b86e3..46ed8282bcd9a 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -280,7 +280,7 @@ protected noncomputable def span : Basis ι R (span R (range v)) := rfl have h₂ : map (Submodule.subtype (span R (range v))) (span R (range fun i => ⟨v i, this i⟩)) = span R (range v) := by - rw [← span_image, Submodule.coeSubtype] + rw [← span_image, Submodule.coe_subtype] -- Porting note: why doesn't `rw [h₁]` work here? exact congr_arg _ h₁ have h₃ : (x : M) ∈ map (Submodule.subtype (span R (range v))) diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index f4d1dcbf680e7..169819f4632df 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -283,7 +283,7 @@ lemma ker_restrict_eq_of_codisjoint {p q : Submodule R M} (hpq : Codisjoint p q) {B : LinearMap.BilinForm R M} (hB : ∀ x ∈ p, ∀ y ∈ q, B x y = 0) : LinearMap.ker (B.restrict p) = (LinearMap.ker B).comap p.subtype := by ext ⟨z, hz⟩ - simp only [LinearMap.mem_ker, Submodule.mem_comap, Submodule.coeSubtype] + simp only [LinearMap.mem_ker, Submodule.mem_comap, Submodule.coe_subtype] refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · ext w obtain ⟨x, hx, y, hy, rfl⟩ := Submodule.exists_add_eq_of_codisjoint hpq w @@ -294,7 +294,7 @@ lemma ker_restrict_eq_of_codisjoint {p q : Submodule R M} (hpq : Codisjoint p q) lemma inf_orthogonal_self_le_ker_restrict {W : Submodule R M} (b₁ : B.IsRefl) : W ⊓ B.orthogonal W ≤ (LinearMap.ker <| B.restrict W).map W.subtype := by rintro v ⟨hv : v ∈ W, hv' : v ∈ B.orthogonal W⟩ - simp only [Submodule.mem_map, mem_ker, restrict_apply, Submodule.coeSubtype, Subtype.exists, + simp only [Submodule.mem_map, mem_ker, restrict_apply, Submodule.coe_subtype, Subtype.exists, exists_and_left, exists_prop, exists_eq_right_right] refine ⟨?_, hv⟩ ext ⟨w, hw⟩ diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean index 1e9ee43c0a118..e06d940aa2396 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean @@ -116,7 +116,7 @@ def ofProd : CliffordAlgebra (Q₁.prod Q₂) →ₐ[R] (evenOdd Q₁ ᵍ⊗[R] ∘ₗ (evenOdd Q₂ 1).subtype ∘ₗ (ι Q₂).codRestrict _ (ι_mem_evenOdd_one Q₂)), fun m => by simp_rw [LinearMap.coprod_apply, LinearMap.coe_comp, Function.comp_apply, - AlgHom.toLinearMap_apply, QuadraticMap.prod_apply, Submodule.coeSubtype, + AlgHom.toLinearMap_apply, QuadraticMap.prod_apply, Submodule.coe_subtype, GradedTensorProduct.includeLeft_apply, GradedTensorProduct.includeRight_apply, map_add, add_mul, mul_add, GradedTensorProduct.algebraMap_def, GradedTensorProduct.tmul_one_mul_one_tmul, GradedTensorProduct.tmul_one_mul_coe_tmul, diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index fde50b6846a2f..7f631f041c8c3 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -299,7 +299,7 @@ theorem biSup_eq_range_dfinsupp_lsum (p : ι → Prop) [DecidablePred p] (S : ι apply le_antisymm · refine iSup₂_le fun i hi y hy => ⟨DFinsupp.single i ⟨y, hy⟩, ?_⟩ rw [LinearMap.comp_apply, filterLinearMap_apply, filter_single_pos _ _ hi] - simp only [lsum_apply_apply, sumAddHom_single, LinearMap.toAddMonoidHom_coe, coeSubtype] + simp only [lsum_apply_apply, sumAddHom_single, LinearMap.toAddMonoidHom_coe, coe_subtype] · rintro x ⟨v, rfl⟩ refine dfinsupp_sumAddHom_mem _ _ _ fun i _ => ?_ refine mem_iSup_of_mem i ?_ @@ -322,7 +322,7 @@ theorem mem_iSup_iff_exists_dfinsupp' (p : ι → Submodule R N) [∀ (i) (x : p (x : N) : x ∈ iSup p ↔ ∃ f : Π₀ i, p i, (f.sum fun i xi => ↑xi) = x := by rw [mem_iSup_iff_exists_dfinsupp] simp_rw [DFinsupp.lsum_apply_apply, DFinsupp.sumAddHom_apply, - LinearMap.toAddMonoidHom_coe, coeSubtype] + LinearMap.toAddMonoidHom_coe, coe_subtype] theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop) [DecidablePred p] (S : ι → Submodule R N) (x : N) : diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index 7830bfb5e4d7f..f2439bd876a5f 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -160,7 +160,7 @@ lemma finrank_maxGenEigenspace (φ : Module.End K M) : apply b.ext simp only [Basis.prod_apply, coe_inl, coe_inr, prodMap_apply, LinearEquiv.conj_apply, LinearEquiv.symm_symm, Submodule.coe_prodEquivOfIsCompl, coe_comp, LinearEquiv.coe_coe, - Function.comp_apply, coprod_apply, Submodule.coeSubtype, map_add, Sum.forall, Sum.elim_inl, + Function.comp_apply, coprod_apply, Submodule.coe_subtype, map_add, Sum.forall, Sum.elim_inl, map_zero, ZeroMemClass.coe_zero, add_zero, LinearEquiv.eq_symm_apply, and_self, Submodule.coe_prodEquivOfIsCompl', restrict_coe_apply, implies_true, Sum.elim_inr, zero_add, e, V, W, ψ, F, G, b] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 7209c69493247..9b95709365e77 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -221,7 +221,7 @@ theorem _root_.Submodule.eq_top_of_finrank_eq [FiniteDimensional K V] {S : Submo (by rw [Set.card_image_of_injective _ Subtype.coe_injective, ← finrank_eq_card_basis bS, ← finrank_eq_card_basis b, h]) - rw [← b.span_eq, b_eq, Basis.coe_extend, Subtype.range_coe, ← this, ← Submodule.coeSubtype, + rw [← b.span_eq, b_eq, Basis.coe_extend, Subtype.range_coe, ← this, ← Submodule.coe_subtype, span_image] have := bS.span_eq rw [bS_eq, Basis.coe_ofVectorSpace, Subtype.range_coe] at this diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index fe713da58c68c..810b8a01afebd 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -365,7 +365,7 @@ theorem supported_iUnion {δ : Type*} (s : δ → Set α) : · exact zero_mem _ · refine fun x a l _ _ => add_mem ?_ by_cases h : ∃ i, x ∈ s i - · simp only [mem_comap, coe_comp, coeSubtype, Function.comp_apply, restrictDom_apply, + · simp only [mem_comap, coe_comp, coe_subtype, Function.comp_apply, restrictDom_apply, mem_iUnion, h, filter_single_of_pos] cases' h with i hi exact le_iSup (fun i => supported M R (s i)) i (single_mem_supported R _ hi) diff --git a/Mathlib/LinearAlgebra/FreeModule/PID.lean b/Mathlib/LinearAlgebra/FreeModule/PID.lean index c5c0346f9a5de..3e5251f600539 100644 --- a/Mathlib/LinearAlgebra/FreeModule/PID.lean +++ b/Mathlib/LinearAlgebra/FreeModule/PID.lean @@ -251,7 +251,7 @@ theorem Submodule.basis_of_pid_aux [Finite ι] {O : Type*} [AddCommGroup O] [Mod refine ⟨-b, Submodule.mem_map.mpr ⟨⟨_, N.sub_mem zN (N.smul_mem b yN)⟩, ?_, ?_⟩⟩ · refine LinearMap.mem_ker.mpr (show ϕ (⟨z, N_le_M zN⟩ - b • ⟨y, N_le_M yN⟩) = 0 from ?_) rw [LinearMap.map_sub, LinearMap.map_smul, hb, ϕy_eq, smul_eq_mul, mul_comm, sub_self] - · simp only [sub_eq_add_neg, neg_smul, coeSubtype] + · simp only [sub_eq_add_neg, neg_smul, coe_subtype] -- And extend a basis for `M'` with `y'` intro m' hn'm' bM' refine ⟨Nat.succ_le_succ hn'm', ?_, ?_⟩ @@ -270,7 +270,7 @@ theorem Submodule.basis_of_pid_aux [Finite ι] {O : Type*} [AddCommGroup O] [Mod · simp only [Fin.cons_zero, Fin.castLE_zero] exact a_smul_y'.symm · rw [Fin.castLE_succ] - simp only [Fin.cons_succ, Function.comp_apply, coe_inclusion, map_coe, coeSubtype, h i] + simp only [Fin.cons_succ, Function.comp_apply, coe_inclusion, map_coe, coe_subtype, h i] /-- A submodule of a free `R`-module of finite rank is also a free `R`-module of finite rank, if `R` is a principal ideal domain. diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 7a9bb31cd43fa..54a39b6372052 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -83,7 +83,7 @@ theorem quotientInfEquivSupQuotient_surjective (p p' : Submodule R M) : rw [← range_eq_top, quotientInfToSupQuotient, range_liftQ, eq_top_iff'] rintro ⟨x, hx⟩; rcases mem_sup.1 hx with ⟨y, hy, z, hz, rfl⟩ use ⟨y, hy⟩; apply (Submodule.Quotient.eq _).2 - simp only [mem_comap, map_sub, coeSubtype, coe_inclusion, sub_add_cancel_left, neg_mem_iff, hz] + simp only [mem_comap, map_sub, coe_subtype, coe_inclusion, sub_add_cancel_left, neg_mem_iff, hz] /-- Second Isomorphism Law : the canonical map from `p/(p ∩ p')` to `(p+p')/p'` as a linear isomorphism. @@ -122,7 +122,7 @@ theorem quotientInfEquivSupQuotient_symm_apply_eq_zero_iff {p p' : Submodule R M (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.coeSubtype] + rw [_root_.map_zero, Quotient.mk_eq_zero, mem_comap, Submodule.coe_subtype] 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/LinearAlgebra/Projection.lean b/Mathlib/LinearAlgebra/Projection.lean index e730906ab556c..d7766cbbf641f 100644 --- a/Mathlib/LinearAlgebra/Projection.lean +++ b/Mathlib/LinearAlgebra/Projection.lean @@ -393,10 +393,10 @@ theorem eq_conj_prod_map' {f : E →ₗ[R] E} (h : IsProj p f) : prodMap id 0 ∘ₗ (p.prodEquivOfIsCompl (ker f) h.isCompl).symm.toLinearMap := by rw [← LinearMap.comp_assoc, LinearEquiv.eq_comp_toLinearMap_symm] ext x - · simp only [coe_prodEquivOfIsCompl, comp_apply, coe_inl, coprod_apply, coeSubtype, + · simp only [coe_prodEquivOfIsCompl, comp_apply, coe_inl, coprod_apply, coe_subtype, _root_.map_zero, add_zero, h.map_id x x.2, prodMap_apply, id_apply] · simp only [coe_prodEquivOfIsCompl, comp_apply, coe_inr, coprod_apply, _root_.map_zero, - coeSubtype, zero_add, map_coe_ker, prodMap_apply, zero_apply, add_zero] + coe_subtype, zero_add, map_coe_ker, prodMap_apply, zero_apply, add_zero] end IsProj diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index 608cd74b64296..14d7396e79842 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -547,7 +547,7 @@ lemma Ideal.map_includeLeft_eq (I : Ideal A) : simp only [map_zero, smul_eq_mul, mul_zero] | tmul x y => use (a • x) ⊗ₜ[R] (b * y) - simp only [LinearMap.lTensor_tmul, Submodule.coeSubtype, smul_eq_mul, tmul_mul_tmul] + simp only [LinearMap.lTensor_tmul, Submodule.coe_subtype, smul_eq_mul, tmul_mul_tmul] with_unfolding_all rfl | add x y hx hy => obtain ⟨x', hx'⟩ := hx @@ -566,7 +566,7 @@ lemma Ideal.map_includeLeft_eq (I : Ideal A) : rw [map_zero] apply zero_mem | tmul a b => - simp only [LinearMap.rTensor_tmul, Submodule.coeSubtype] + simp only [LinearMap.rTensor_tmul, Submodule.coe_subtype] suffices (a : A) ⊗ₜ[R] b = ((1 : A) ⊗ₜ[R] b) * ((a : A) ⊗ₜ[R] (1 : B)) by simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, Submodule.mem_toAddSubmonoid, Submodule.restrictScalars_mem] @@ -615,7 +615,7 @@ lemma Ideal.map_includeRight_eq (I : Ideal B) : simp only [map_zero, smul_eq_mul, mul_zero] | tmul x y => use (a * x) ⊗ₜ[R] (b •y) - simp only [LinearMap.lTensor_tmul, Submodule.coeSubtype, smul_eq_mul, tmul_mul_tmul] + simp only [LinearMap.lTensor_tmul, Submodule.coe_subtype, smul_eq_mul, tmul_mul_tmul] rfl | add x y hx hy => obtain ⟨x', hx'⟩ := hx @@ -634,7 +634,7 @@ lemma Ideal.map_includeRight_eq (I : Ideal B) : rw [map_zero] apply zero_mem | tmul a b => - simp only [LinearMap.lTensor_tmul, Submodule.coeSubtype] + simp only [LinearMap.lTensor_tmul, Submodule.coe_subtype] suffices a ⊗ₜ[R] (b : B) = (a ⊗ₜ[R] (1 : B)) * ((1 : A) ⊗ₜ[R] (b : B)) by rw [this] simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, diff --git a/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean b/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean index dbf016581b5a2..8a6dcf8f11f00 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean @@ -146,7 +146,7 @@ theorem vanishesTrivially_of_sum_tmul_eq_zero (hm : Submodule.span R (Set.range symm at hkn simp only [map_sum, finsuppScalarLeft_apply_tmul, zero_smul, Finsupp.single_zero, Finsupp.sum_single_index, one_smul, Finsupp.finset_sum_apply, Finsupp.single_apply, - Finset.sum_ite_eq', Finset.mem_univ, ↓reduceIte, rTensor_tmul, coeSubtype, Finsupp.sum_apply, + Finset.sum_ite_eq', Finset.mem_univ, ↓reduceIte, rTensor_tmul, coe_subtype, Finsupp.sum_apply, Finsupp.sum_ite_eq', Finsupp.mem_support_iff, ne_eq, ite_not, en] at hkn simp only [Finset.univ_eq_attach, Finset.sum_attach ma (fun x ↦ (x.1 : ι →₀ R) i • x.2)] convert hkn using 2 with x _ @@ -181,11 +181,11 @@ theorem vanishesTrivially_of_sum_tmul_eq_zero_of_rTensor_injective set m' : ι → span R (Set.range m) := Subtype.coind m mem_M' with m'_eq have hm' : span R (Set.range m') = ⊤ := by apply map_injective_of_injective (injective_subtype (span R (Set.range m))) - rw [Submodule.map_span, Submodule.map_top, range_subtype, coeSubtype, ← Set.range_comp] + rw [Submodule.map_span, Submodule.map_top, range_subtype, coe_subtype, ← Set.range_comp] rfl have hm'n : ∑ i, m' i ⊗ₜ n i = (0 : span R (Set.range m) ⊗[R] N) := by apply hm - simp only [m'_eq, map_sum, rTensor_tmul, coeSubtype, Subtype.coind_coe, _root_.map_zero, hmn] + simp only [m'_eq, map_sum, rTensor_tmul, coe_subtype, Subtype.coind_coe, _root_.map_zero, hmn] have : VanishesTrivially R m' n := vanishesTrivially_of_sum_tmul_eq_zero R hm' hm'n unfold VanishesTrivially at this ⊢ convert this with κ _ a y j @@ -218,7 +218,7 @@ theorem rTensor_injective_of_forall_vanishesTrivially obtain ⟨s, rfl⟩ := exists_finset x rw [← Finset.sum_attach] apply sum_tmul_eq_zero_of_vanishesTrivially - simp only [map_sum, rTensor_tmul, coeSubtype] at hx + simp only [map_sum, rTensor_tmul, coe_subtype] at hx have := hMN ((Finset.sum_attach s _).trans hx) unfold VanishesTrivially at this ⊢ convert this with κ _ a y j diff --git a/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean b/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean index 3d5cc80dab357..47f93f496cc3b 100644 --- a/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean +++ b/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean @@ -61,7 +61,7 @@ instance (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : · refine Submonoid.mul_mem _ hd (mem_nonZeroDivisors_of_ne_zero ?_) rw [Nat.cast_ne_zero, ne_eq, Ideal.absNorm_eq_zero_iff] exact FractionalIdeal.num_eq_zero_iff.not.mpr <| Units.ne_zero I - · simp_rw [LinearMap.coe_restrictScalars, Submodule.coeSubtype] at h ⊢ + · simp_rw [LinearMap.coe_restrictScalars, Submodule.coe_subtype] at h ⊢ rw [← h] simp only [Submonoid.mk_smul, zsmul_eq_mul, Int.cast_mul, Int.cast_natCast, algebraMap_int_eq, eq_intCast, map_intCast] diff --git a/Mathlib/Probability/ConditionalExpectation.lean b/Mathlib/Probability/ConditionalExpectation.lean index 66719c1935974..18a5ab72ec167 100644 --- a/Mathlib/Probability/ConditionalExpectation.lean +++ b/Mathlib/Probability/ConditionalExpectation.lean @@ -58,11 +58,11 @@ theorem condexp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinit · have heq₁ : (fun f : lpMeas E ℝ m₁ 1 μ => ∫ x, (f : Ω → E) x ∂μ) = (fun f : Lp E 1 μ => ∫ x, f x ∂μ) ∘ Submodule.subtypeL _ := by refine funext fun f => integral_congr_ae ?_ - simp_rw [Submodule.coe_subtypeL', Submodule.coeSubtype]; norm_cast + simp_rw [Submodule.coe_subtypeL', Submodule.coe_subtype]; norm_cast have heq₂ : (fun f : lpMeas E ℝ m₁ 1 μ => ∫ x in s, (f : Ω → E) x ∂μ) = (fun f : Lp E 1 μ => ∫ x in s, f x ∂μ) ∘ Submodule.subtypeL _ := by refine funext fun f => integral_congr_ae (ae_restrict_of_ae ?_) - simp_rw [Submodule.coe_subtypeL', Submodule.coeSubtype] + simp_rw [Submodule.coe_subtypeL', Submodule.coe_subtype] exact Eventually.of_forall fun _ => (by trivial) refine isClosed_eq (Continuous.const_smul ?_ _) ?_ · rw [heq₁] diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 7592d42a7c793..813b32f84d0ac 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -172,7 +172,7 @@ instance directSum (ι : Type v) (M : ι → Type w) [(i : ι) → AddCommGroup apply TensorProduct.ext' intro a m simp only [ρ, ψ, φ, η, η₁, coe_comp, LinearEquiv.coe_coe, Function.comp_apply, - directSumRight_symm_lof_tmul, rTensor_tmul, Submodule.coeSubtype, lid_tmul, map_smul] + directSumRight_symm_lof_tmul, rTensor_tmul, Submodule.coe_subtype, lid_tmul, map_smul] rw [DirectSum.component.of, DirectSum.component.of] by_cases h₂ : j = i · subst j; simp From d377e29420861c72b9298587f37a322715ea4f01 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 28 Sep 2024 17:13:44 +0000 Subject: [PATCH 017/174] chore: update Mathlib dependencies 2024-09-28 (#17226) 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 c54010edda52c..83f9c543ecef0 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "e0b13c946e9c3805f1eec785c72955e103a9cbaf", + "rev": "51c38e3828d06c82741a7a65df93611e2ce209e1", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 332b6cd7794f09d3b74b310d159f1735ac6e7ffa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 17:13:45 +0000 Subject: [PATCH 018/174] chore(Finset/Density): Fix statement and make lemmas simp (#17227) From LeanAPAP --- Mathlib/Data/Finset/Density.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index 43125f7fcc4f3..002513c40e808 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -110,23 +110,23 @@ lemma dens_image [Fintype β] [DecidableEq β] {f : α → β} (hf : Bijective f (s.image f).dens = s.dens := by simpa [map_eq_image, -dens_map_equiv] using dens_map_equiv (.ofBijective f hf) -lemma card_mul_dens (s : Finset α) : Fintype.card α * s.dens = s.card := by +@[simp] lemma card_mul_dens (s : Finset α) : Fintype.card α * s.dens = s.card := by cases isEmpty_or_nonempty α · simp [Subsingleton.elim s ∅] rw [dens, mul_div_cancel₀] exact mod_cast Fintype.card_ne_zero -lemma dens_mul_card (s : Finset α) : s.dens * Fintype.card α = s.card := by +@[simp] lemma dens_mul_card (s : Finset α) : s.dens * Fintype.card α = s.card := by rw [mul_comm, card_mul_dens] section Semifield variable [Semifield 𝕜] [CharZero 𝕜] -lemma natCast_card_mul_nnratCast_dens (s : Finset α) : (Fintype.card α * s.dens : 𝕜) = s.card := - mod_cast s.card_mul_dens +@[simp] lemma natCast_card_mul_nnratCast_dens (s : Finset α) : + (Fintype.card α * s.dens : 𝕜) = s.card := mod_cast s.card_mul_dens -lemma nnratCast_dens_mul_natCast_card (s : Finset α) : s.dens * Fintype.card α = s.card := - mod_cast s.dens_mul_card +@[simp] lemma nnratCast_dens_mul_natCast_card (s : Finset α) : + (s.dens * Fintype.card α : 𝕜) = s.card := mod_cast s.dens_mul_card @[norm_cast] lemma nnratCast_dens (s : Finset α) : (s.dens : 𝕜) = s.card / Fintype.card α := by simp [dens] From a087cde0a3dc36762e4954ff1b2763540431aa18 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 18:08:36 +0000 Subject: [PATCH 019/174] chore(Num): Undeprecate (#17207) Jireh deprecated these declarations in #607 because they used `bit0` and `bit1`. They since have been changed to use `x + x` and `x + x + 1`, but were not un-deprecated. They are used in mathlib. --- Mathlib/Data/Num/Basic.lean | 18 +++++------------- 1 file changed, 5 insertions(+), 13 deletions(-) diff --git a/Mathlib/Data/Num/Basic.lean b/Mathlib/Data/Num/Basic.lean index 1353fdd100589..5e35ca5550ad8 100644 --- a/Mathlib/Data/Num/Basic.lean +++ b/Mathlib/Data/Num/Basic.lean @@ -169,33 +169,27 @@ section variable {α : Type*} [One α] [Add α] -section deprecated -set_option linter.deprecated false - /-- `castPosNum` casts a `PosNum` into any type which has `1` and `+`. -/ -@[deprecated (since := "2022-11-18"), coe] +@[coe] def castPosNum : PosNum → α | 1 => 1 | PosNum.bit0 a => castPosNum a + castPosNum a | PosNum.bit1 a => castPosNum a + castPosNum a + 1 /-- `castNum` casts a `Num` into any type which has `0`, `1` and `+`. -/ -@[deprecated (since := "2022-11-18"), coe] +@[coe] def castNum [Zero α] : Num → α | 0 => 0 | Num.pos p => castPosNum p -- see Note [coercion into rings] -@[deprecated (since := "2023-03-31")] instance (priority := 900) posNumCoe : CoeHTCT PosNum α := +instance (priority := 900) posNumCoe : CoeHTCT PosNum α := ⟨castPosNum⟩ -- see Note [coercion into rings] -@[deprecated (since := "2023-03-31")] instance (priority := 900) numNatCoe [Zero α] : CoeHTCT Num α := ⟨castNum⟩ -end deprecated - instance : Repr PosNum := ⟨fun n _ => repr (n : ℕ)⟩ @@ -593,19 +587,17 @@ def gcd (a b : ZNum) : Num := end ZNum section - -set_option linter.deprecated false variable {α : Type*} [Zero α] [One α] [Add α] [Neg α] /-- `castZNum` casts a `ZNum` into any type which has `0`, `1`, `+` and `neg` -/ -@[deprecated (since := "2022-11-18"), coe] +@[coe] def castZNum : ZNum → α | 0 => 0 | ZNum.pos p => p | ZNum.neg p => -p -- see Note [coercion into rings] -@[deprecated (since := "2023-03-31")] instance (priority := 900) znumCoe : CoeHTCT ZNum α := +instance (priority := 900) znumCoe : CoeHTCT ZNum α := ⟨castZNum⟩ instance : Repr ZNum := From 5a5f390c3c6fcea343f87026bac4cc6c0d40d35f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 19:13:00 +0000 Subject: [PATCH 020/174] feat(AddChar): more basic lemmas (#17018) Also unsimp `AddChar.div_apply`/`AddChar.sub_apply` because it doesn't do the right thing in my application. From LeanAPAP --- Mathlib/Algebra/Group/AddChar.lean | 31 +++++++++++++++++++++++++----- 1 file changed, 26 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 28c90b17cfdf1..f9a1d383c185c 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -195,6 +195,16 @@ instance instZero : Zero (AddChar A M) := ⟨1⟩ lemma one_eq_zero : (1 : AddChar A M) = (0 : AddChar A M) := rfl +@[simp, norm_cast] lemma coe_eq_one : ⇑ψ = 1 ↔ ψ = 0 := by rw [← coe_zero, DFunLike.coe_fn_eq] + +@[simp] lemma toMonoidHomEquiv_zero : toMonoidHomEquiv (0 : AddChar A M) = 1 := rfl +@[simp] lemma toMonoidHomEquiv_symm_one : + toMonoidHomEquiv.symm (1 : Multiplicative A →* M) = 0 := rfl + +@[simp] lemma toAddMonoidHomEquiv_zero : toAddMonoidHomEquiv (0 : AddChar A M) = 0 := rfl +@[simp] lemma toAddMonoidHomEquiv_symm_zero : + toAddMonoidHomEquiv.symm (0 : A →+ Additive M) = 0 := rfl + instance instInhabited : Inhabited (AddChar A M) := ⟨1⟩ /-- Composing a `MonoidHom` with an `AddChar` yields another `AddChar`. -/ @@ -252,6 +262,8 @@ set_option linter.deprecated false in lemma isNontrivial_iff_ne_trivial (ψ : AddChar A M) : IsNontrivial ψ ↔ ψ ≠ 1 := not_forall.symm.trans (DFunLike.ext_iff (f := ψ) (g := 1)).symm.not +noncomputable instance : DecidableEq (AddChar A M) := Classical.decEq _ + end Basic section toCommMonoid @@ -291,6 +303,11 @@ lemma mul_eq_add (ψ χ : AddChar A M) : ψ * χ = ψ + χ := rfl lemma pow_eq_nsmul (ψ : AddChar A M) (n : ℕ) : ψ ^ n = n • ψ := rfl lemma prod_eq_sum (s : Finset ι) (ψ : ι → AddChar A M) : ∏ i in s, ψ i = ∑ i in s, ψ i := rfl +@[simp] lemma toMonoidHomEquiv_add (ψ φ : AddChar A M) : + toMonoidHomEquiv (ψ + φ) = toMonoidHomEquiv ψ * toMonoidHomEquiv φ := rfl +@[simp] lemma toMonoidHomEquiv_symm_mul (ψ φ : Multiplicative A →* M) : + toMonoidHomEquiv.symm (ψ * φ) = toMonoidHomEquiv.symm ψ + toMonoidHomEquiv.symm φ := rfl + /-- The natural equivalence to `(Multiplicative A →* M)` is a monoid isomorphism. -/ def toMonoidHomMulEquiv : AddChar A M ≃* (Multiplicative A →* M) := { toMonoidHomEquiv with map_mul' := fun φ ψ ↦ by rfl } @@ -356,8 +373,8 @@ instance : AddCommGroup (AddChar A M) := Additive.addCommGroup @[simp] lemma inv_apply (ψ : AddChar A M) (a : A) : ψ⁻¹ a = ψ (-a) := rfl @[simp] lemma neg_apply (ψ : AddChar A M) (a : A) : (-ψ) a = ψ (-a) := rfl -@[simp] lemma div_apply (ψ χ : AddChar A M) (a : A) : (ψ / χ) a = ψ a * χ (-a) := rfl -@[simp] lemma sub_apply (ψ χ : AddChar A M) (a : A) : (ψ - χ) a = ψ a * χ (-a) := rfl +lemma div_apply (ψ χ : AddChar A M) (a : A) : (ψ / χ) a = ψ a * χ (-a) := rfl +lemma sub_apply (ψ χ : AddChar A M) (a : A) : (ψ - χ) a = ψ a * χ (-a) := rfl end fromAddCommGroup @@ -387,8 +404,7 @@ lemma map_zsmul_eq_zpow (ψ : AddChar A M) (n : ℤ) (a : A) : ψ (n • a) = ( end fromAddGrouptoDivisionMonoid -section fromAddGrouptoDivisionCommMonoid - +section fromAddCommGrouptoDivisionCommMonoid variable {A M : Type*} [AddCommGroup A] [DivisionCommMonoid M] lemma inv_apply' (ψ : AddChar A M) (a : A) : ψ⁻¹ a = (ψ a)⁻¹ := by rw [inv_apply, map_neg_eq_inv] @@ -400,13 +416,18 @@ lemma div_apply' (ψ χ : AddChar A M) (a : A) : (ψ / χ) a = ψ a / χ a := by lemma sub_apply' (ψ χ : AddChar A M) (a : A) : (ψ - χ) a = ψ a / χ a := by rw [sub_apply, map_neg_eq_inv, div_eq_mul_inv] +@[simp] lemma zsmul_apply (n : ℤ) (ψ : AddChar A M) (a : A) : (n • ψ) a = ψ a ^ n := by + cases n <;> simp [-neg_apply, neg_apply'] + +@[simp] lemma zpow_apply (ψ : AddChar A M) (n : ℤ) (a : A) : (ψ ^ n) a = ψ a ^ n := zsmul_apply .. + lemma map_sub_eq_div (ψ : AddChar A M) (a b : A) : ψ (a - b) = ψ a / ψ b := ψ.toMonoidHom.map_div _ _ lemma injective_iff {ψ : AddChar A M} : Injective ψ ↔ ∀ ⦃x⦄, ψ x = 1 → x = 0 := ψ.toMonoidHom.ker_eq_bot_iff.symm.trans eq_bot_iff -end fromAddGrouptoDivisionCommMonoid +end fromAddCommGrouptoDivisionCommMonoid section MonoidWithZero variable {A M₀ : Type*} [AddGroup A] [MonoidWithZero M₀] [Nontrivial M₀] From 4ceb2ea935085e8145e880f754cef760063f16de Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 19:40:23 +0000 Subject: [PATCH 021/174] feat: order properties of `Finset.expect` (#16946) Also add a few basic lemmas and fix the delaborator in `Algebra.BigOperators.Prereqs`. From LeanAPAP --- Mathlib.lean | 1 + Mathlib/Algebra/BigOperators/Expect.lean | 26 ++- .../Algebra/Order/BigOperators/Expect.lean | 211 ++++++++++++++++++ Mathlib/Analysis/MeanInequalities.lean | 14 ++ 4 files changed, 250 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Algebra/Order/BigOperators/Expect.lean diff --git a/Mathlib.lean b/Mathlib.lean index 94bbbdbd563f3..6894a6f609343 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -546,6 +546,7 @@ import Mathlib.Algebra.Order.Antidiag.Prod import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Archimedean.Hom import Mathlib.Algebra.Order.Archimedean.Submonoid +import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Algebra.Order.BigOperators.Group.Multiset diff --git a/Mathlib/Algebra/BigOperators/Expect.lean b/Mathlib/Algebra/BigOperators/Expect.lean index b64a24d549e6c..cf6a0e4d72664 100644 --- a/Mathlib/Algebra/BigOperators/Expect.lean +++ b/Mathlib/Algebra/BigOperators/Expect.lean @@ -5,8 +5,11 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Algebra.Rat import Mathlib.Algebra.BigOperators.GroupWithZero.Action +import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Module.Pi +import Mathlib.Data.Finset.Density import Mathlib.Data.Fintype.BigOperators /-! @@ -37,7 +40,7 @@ combination operator. ## TODO * Connect `Finset.expect` with the expectation over `s` in the probability theory sense. -* Give a formulation of Jensen's inequality and the Cauchy-Schwarz inequality in this language. +* Give a formulation of Jensen's inequality in this language. -/ open Finset Function @@ -98,7 +101,7 @@ to show the domain type when the expect is over `Finset.univ`. -/ `(bigOpBinder| $(.mk i):ident) `(𝔼 $binder:bigOpBinder, $body) else - let ss ← withNaryArg 3 <| delab + let ss ← withNaryArg 4 <| delab `(𝔼 $(.mk i):ident ∈ $ss, $body) end BigOperators @@ -159,6 +162,12 @@ lemma expect_ite_zero (s : Finset ι) (p : ι → Prop) [DecidablePred p] 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 + 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 split_ifs <;> simp [expect, *] @@ -358,6 +367,11 @@ lemma expect_div (s : Finset ι) (f : ι → M) (a : M) : (𝔼 i ∈ s, f i) / simp_rw [div_eq_mul_inv, expect_mul] end Semifield + +@[simp] lemma expect_apply {α : Type*} {π : α → Type*} [∀ a, CommSemiring (π a)] + [∀ a, Module ℚ≥0 (π a)] (s : Finset ι) (f : ι → ∀ a, π a) (a : α) : + (𝔼 i ∈ s, f i) a = 𝔼 i ∈ s, f i a := by simp [expect] + end Finset namespace algebraMap @@ -399,6 +413,10 @@ lemma expect_ite_zero (p : ι → Prop) [DecidablePred p] (h : ∀ i j, p i → variable [DecidableEq ι] +@[simp] lemma expect_ite_mem (s : Finset ι) (f : ι → M) : + 𝔼 i, (if i ∈ s then f i else 0) = s.dens • 𝔼 i ∈ s, f i := by + simp [Finset.expect_ite_mem, dens] + lemma expect_dite_eq (i : ι) (f : ∀ j, i = j → M) : 𝔼 j, (if h : i = j then f j h else 0) = f i rfl /ℚ card ι := by simp [card_univ] @@ -418,5 +436,9 @@ variable [Semiring M] [Module ℚ≥0 M] lemma expect_one [Nonempty ι] : 𝔼 _i : ι, (1 : M) = 1 := expect_const _ +lemma expect_mul_expect [IsScalarTower ℚ≥0 M M] [SMulCommClass ℚ≥0 M M] (f : ι → M) + (g : κ → M) : (𝔼 i, f i) * 𝔼 j, g j = 𝔼 i, 𝔼 j, f i * g j := + Finset.expect_mul_expect .. + end Semiring end Fintype diff --git a/Mathlib/Algebra/Order/BigOperators/Expect.lean b/Mathlib/Algebra/Order/BigOperators/Expect.lean new file mode 100644 index 0000000000000..30064a4bcb627 --- /dev/null +++ b/Mathlib/Algebra/Order/BigOperators/Expect.lean @@ -0,0 +1,211 @@ +/- +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.BigOperators.Expect +import Mathlib.Algebra.Module.Rat +import Mathlib.Algebra.Order.BigOperators.Ring.Finset +import Mathlib.Algebra.Order.Module.Rat + +/-! +# Order properties of the average over a finset +-/ + +open Function +open Fintype (card) +open scoped BigOperators Pointwise NNRat + +variable {ι κ α β R : Type*} + +local notation a " /ℚ " q => (q : ℚ≥0)⁻¹ • a + +namespace Finset +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid α] [Module ℚ≥0 α] [OrderedAddCommMonoid β] [Module ℚ≥0 β] + {s : Finset ι} {f g : ι → α} + +lemma expect_eq_zero_iff_of_nonneg (hs : s.Nonempty) (hf : ∀ i ∈ s, 0 ≤ f i) : + 𝔼 i ∈ s, f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + simp [expect, sum_eq_zero_iff_of_nonneg hf, hs.ne_empty] + +lemma expect_eq_zero_iff_of_nonpos (hs : s.Nonempty) (hf : ∀ i ∈ s, f i ≤ 0) : + 𝔼 i ∈ s, f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + simp [expect, sum_eq_zero_iff_of_nonpos hf, hs.ne_empty] + +section PosSMulMono +variable [PosSMulMono ℚ≥0 α] {a : α} + +lemma expect_le_expect (hfg : ∀ i ∈ s, f i ≤ g i) : 𝔼 i ∈ s, f i ≤ 𝔼 i ∈ s, g i := + smul_le_smul_of_nonneg_left (sum_le_sum hfg) <| by positivity + +/-- This is a (beta-reduced) version of the standard lemma `Finset.expect_le_expect`, +convenient for the `gcongr` tactic. -/ +@[gcongr] +lemma _root_.GCongr.expect_le_expect (h : ∀ i ∈ s, f i ≤ g i) : s.expect f ≤ s.expect g := + Finset.expect_le_expect h + +lemma expect_le (hs : s.Nonempty) (h : ∀ x ∈ s, f x ≤ a) : 𝔼 i ∈ s, f i ≤ a := + (inv_smul_le_iff_of_pos <| mod_cast hs.card_pos).2 <| by + rw [Nat.cast_smul_eq_nsmul]; exact sum_le_card_nsmul _ _ _ h + +lemma le_expect (hs : s.Nonempty) (h : ∀ x ∈ s, a ≤ f x) : a ≤ 𝔼 i ∈ s, f i := + (le_inv_smul_iff_of_pos <| mod_cast hs.card_pos).2 <| by + rw [Nat.cast_smul_eq_nsmul]; exact card_nsmul_le_sum _ _ _ h + +lemma expect_nonneg (hf : ∀ i ∈ s, 0 ≤ f i) : 0 ≤ 𝔼 i ∈ s, f i := + smul_nonneg (by positivity) <| sum_nonneg hf + +end PosSMulMono + +section PosSMulMono +variable {M N : Type*} [AddCommMonoid M] [Module ℚ≥0 M] [OrderedAddCommMonoid N] [Module ℚ≥0 N] + [PosSMulMono ℚ≥0 N] {m : M → N} {p : M → Prop} {f : ι → M} {s : Finset ι} + +/-- Let `{a | p a}` be an additive subsemigroup of an additive commutative monoid `M`. If `m` is a +subadditive function (`m (a + b) ≤ m a + m b`) preserved under division by a natural, `f` is a +function valued in that subsemigroup and `s` is a nonempty set, then +`m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i)`. -/ +lemma le_expect_nonempty_of_subadditive_on_pred (h_add : ∀ a b, p a → p b → m (a + b) ≤ m a + m b) + (hp_add : ∀ a b, p a → p b → p (a + b)) (h_div : ∀ (n : ℕ) a, p a → m (a /ℚ n) = m a /ℚ n) + (hs_nonempty : s.Nonempty) (hs : ∀ i ∈ s, p (f i)) : + m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i) := by + simp only [expect, h_div _ _ (sum_induction_nonempty _ _ hp_add hs_nonempty hs)] + exact smul_le_smul_of_nonneg_left + (le_sum_nonempty_of_subadditive_on_pred _ _ h_add hp_add _ _ hs_nonempty hs) <| by positivity + +/-- If `m : M → N` is a subadditive function (`m (a + b) ≤ m a + m b`) and `s` is a nonempty set, +then `m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i)`. -/ +lemma le_expect_nonempty_of_subadditive (m : M → N) (h_mul : ∀ a b, m (a + b) ≤ m a + m b) + (h_div : ∀ (n : ℕ) a, m (a /ℚ n) = m a /ℚ n) (hs : s.Nonempty) : + m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i) := + le_expect_nonempty_of_subadditive_on_pred (p := fun _ ↦ True) (by simpa) (by simp) (by simpa) hs + (by simp) + +/-- Let `{a | p a}` be a subsemigroup of a commutative monoid `M`. If `m` is a subadditive function +(`m (x + y) ≤ m x + m y`, `m 0 = 0`) preserved under division by a natural and `f` is a function +valued in that subsemigroup, then `m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i)`. -/ +lemma le_expect_of_subadditive_on_pred (h_zero : m 0 = 0) + (h_add : ∀ a b, p a → p b → m (a + b) ≤ m a + m b) (hp_add : ∀ a b, p a → p b → p (a + b)) + (h_div : ∀ (n : ℕ) a, p a → m (a /ℚ n) = m a /ℚ n) + (hs : ∀ i ∈ s, p (f i)) : m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i) := by + obtain rfl | hs_nonempty := s.eq_empty_or_nonempty + · simp [h_zero] + · exact le_expect_nonempty_of_subadditive_on_pred h_add hp_add h_div hs_nonempty hs + +-- TODO: Contribute back better docstring to `le_prod_of_submultiplicative` +/-- If `m` is a subadditive function (`m (x + y) ≤ m x + m y`, `m 0 = 0`) preserved under division +by a natural, then `m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i)`. -/ +lemma le_expect_of_subadditive (h_zero : m 0 = 0) (h_add : ∀ a b, m (a + b) ≤ m a + m b) + (h_div : ∀ (n : ℕ) a, m (a /ℚ n) = m a /ℚ n) : m (𝔼 i ∈ s, f i) ≤ 𝔼 i ∈ s, m (f i) := + le_expect_of_subadditive_on_pred (p := fun _ ↦ True) h_zero (by simpa) (by simp) (by simpa) + (by simp) + +end PosSMulMono +end OrderedAddCommMonoid + +section OrderedCancelAddCommMonoid +variable [OrderedCancelAddCommMonoid α] [Module ℚ≥0 α] {s : Finset ι} {f g : ι → α} +section PosSMulStrictMono +variable [PosSMulStrictMono ℚ≥0 α] + +lemma expect_pos (hf : ∀ i ∈ s, 0 < f i) (hs : s.Nonempty) : 0 < 𝔼 i ∈ s, f i := + smul_pos (inv_pos.2 <| mod_cast hs.card_pos) <| sum_pos hf hs + +end PosSMulStrictMono +end OrderedCancelAddCommMonoid + +section LinearOrderedAddCommMonoid +variable [LinearOrderedAddCommMonoid α] [Module ℚ≥0 α] [PosSMulMono ℚ≥0 α] {s : Finset ι} + {f : ι → α} {a : α} + +lemma exists_lt_of_lt_expect (hs : s.Nonempty) (h : a < 𝔼 i ∈ s, f i) : ∃ x ∈ s, a < f x := by + contrapose! h; exact expect_le hs h + +lemma exists_lt_of_expect_lt (hs : s.Nonempty) (h : 𝔼 i ∈ s, f i < a) : ∃ x ∈ s, f x < a := by + contrapose! h; exact le_expect hs h + +end LinearOrderedAddCommMonoid + +section LinearOrderedAddCommGroup +variable [LinearOrderedAddCommGroup α] [Module ℚ≥0 α] [PosSMulMono ℚ≥0 α] + +-- TODO: Norm version +lemma abs_expect_le_expect_abs (s : Finset ι) (f : ι → α) : |𝔼 i ∈ s, f i| ≤ 𝔼 i ∈ s, |f i| := + le_expect_of_subadditive abs_zero abs_add (fun _ ↦ abs_nnqsmul _) + +end LinearOrderedAddCommGroup + +section LinearOrderedCommSemiring +variable [LinearOrderedCommSemiring R] [ExistsAddOfLE R] [Module ℚ≥0 R] [PosSMulMono ℚ≥0 R] + +/-- **Cauchy-Schwarz inequality** in terms of `Finset.expect`. -/ +lemma expect_mul_sq_le_sq_mul_sq (s : Finset ι) (f g : ι → R) : + (𝔼 i ∈ s, f i * g i) ^ 2 ≤ (𝔼 i ∈ s, f i ^ 2) * 𝔼 i ∈ s, g i ^ 2 := by + simp only [expect, smul_pow, inv_pow, smul_mul_smul_comm, ← sq] + gcongr + exact sum_mul_sq_le_sq_mul_sq .. + +end LinearOrderedCommSemiring +end Finset + +open Finset + +namespace Fintype +variable [Fintype ι] [Fintype κ] + +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid α] [Module ℚ≥0 α] {f : ι → α} + +lemma expect_eq_zero_iff_of_nonneg [Nonempty ι] (hf : 0 ≤ f) : 𝔼 i, f i = 0 ↔ f = 0 := by + simp [expect, sum_eq_zero_iff_of_nonneg hf, univ_nonempty.ne_empty] + +lemma expect_eq_zero_iff_of_nonpos [Nonempty ι] (hf : f ≤ 0) : 𝔼 i, f i = 0 ↔ f = 0 := by + simp [expect, sum_eq_zero_iff_of_nonpos hf, univ_nonempty.ne_empty] + +end OrderedAddCommMonoid +end Fintype + +open Finset + +namespace Mathlib.Meta.Positivity +open Qq Lean Meta Finset +open scoped BigOperators + +/-- Positivity extension for `Finset.expect`. -/ +@[positivity Finset.expect _ _] +def evalFinsetExpect : PositivityExt where eval {u α} zα pα e := do + match e with + | ~q(@Finset.expect $ι _ $instα $instmod $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 + let .some instαordsmul ← trySynthInstanceQ q(PosSMulStrictMono ℚ≥0 $α) | pure none + assumeInstancesCommute + let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody + return some q(@expect_pos $ι $α $pα' $instmod $s $f $instαordsmul (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 instαordmon ← synthInstanceQ q(OrderedAddCommMonoid $α) + let instαordsmul ← synthInstanceQ q(PosSMulMono ℚ≥0 $α) + assumeInstancesCommute + return .nonnegative q(@expect_nonneg $ι $α $instαordmon $instmod $s $f $instαordsmul + fun i _ ↦ $pr i) + | _ => throwError "not Finset.expect" + +example (n : ℕ) (a : ℕ → ℚ) : 0 ≤ 𝔼 j ∈ range n, a j^2 := by positivity +example (a : ULift.{2} ℕ → ℚ) (s : Finset (ULift.{2} ℕ)) : 0 ≤ 𝔼 j ∈ s, a j^2 := by positivity +example (n : ℕ) (a : ℕ → ℚ) : 0 ≤ 𝔼 j : Fin 8, 𝔼 i ∈ range n, (a j^2 + i ^ 2) := by positivity +example (n : ℕ) (a : ℕ → ℚ) : 0 < 𝔼 j : Fin (n + 1), (a j^2 + 1) := by positivity +example (a : ℕ → ℚ) : 0 < 𝔼 j ∈ ({1} : Finset ℕ), (a j^2 + 1) := by positivity + +end Mathlib.Meta.Positivity diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 2426536d2bcae..6f6a2a5350282 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -3,6 +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, Sébastien Gouëzel, Rémy Degenne -/ +import Mathlib.Algebra.BigOperators.Expect import Mathlib.Analysis.Convex.Jensen import Mathlib.Analysis.Convex.SpecificFunctions.Basic import Mathlib.Analysis.SpecialFunctions.Pow.NNReal @@ -105,6 +106,7 @@ less than or equal to the sum of the maximum values of the summands. universe u v open Finset NNReal ENNReal +open scoped BigOperators noncomputable section @@ -660,6 +662,18 @@ lemma inner_le_weight_mul_Lp_of_nonneg (s : Finset ι) {p : ℝ} (hp : 1 ≤ p) norm_cast at * exact NNReal.inner_le_weight_mul_Lp _ hp _ _ +/-- **Weighted Hölder inequality** in terms of `Finset.expect`. -/ +lemma compact_inner_le_weight_mul_Lp_of_nonneg (s : Finset ι) {p : ℝ} (hp : 1 ≤ p) {w f : ι → ℝ} + (hw : ∀ i, 0 ≤ w i) (hf : ∀ i, 0 ≤ f i) : + 𝔼 i ∈ s, w i * f i ≤ (𝔼 i ∈ s, w i) ^ (1 - p⁻¹) * (𝔼 i ∈ s, w i * f i ^ p) ^ p⁻¹ := by + simp_rw [expect_eq_sum_div_card] + rw [div_rpow, div_rpow, div_mul_div_comm, ← rpow_add', sub_add_cancel, rpow_one] + gcongr + · exact inner_le_weight_mul_Lp_of_nonneg s hp _ _ hw hf + any_goals simp + · exact sum_nonneg fun i _ ↦ by have := hw i; have := hf i; positivity + · exact sum_nonneg fun i _ ↦ by have := hw i; positivity + /-- **Hölder inequality**: the scalar product of two functions is bounded by the product of their `L^p` and `L^q` norms when `p` and `q` are conjugate exponents. A version for `ℝ`-valued functions. For an alternative version, convenient if the infinite sums are already expressed as `p`-th powers, From 9b9053eacfc271ea559687f4336b6f3f34d92ecc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 20:52:41 +0000 Subject: [PATCH 022/174] feat(Pointwise): `gcongr` attributes and a few more lemmas (#17233) Also mark more lemmas as `gcongr` From LeanCamCombi --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 22 ++++++++++++++----- Mathlib/Algebra/Group/Pointwise/Set.lean | 20 ++++++++++++----- .../Measure/Lebesgue/EqHaar.lean | 1 - 3 files changed, 30 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index a44299063bf7c..521a3b832afdc 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -221,7 +221,7 @@ attribute [aesop safe apply (rule_sets := [finsetNonempty])] Nonempty.inv Nonemp @[to_additive (attr := simp)] theorem inv_eq_empty : s⁻¹ = ∅ ↔ s = ∅ := image_eq_empty -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem inv_subset_inv (h : s ⊆ t) : s⁻¹ ⊆ t⁻¹ := image_subset_image h @@ -373,7 +373,7 @@ theorem singleton_mul (a : α) : {a} * s = s.image (a * ·) := theorem singleton_mul_singleton (a b : α) : ({a} : Finset α) * {b} = {a * b} := image₂_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem mul_subset_mul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ * t₁ ⊆ s₂ * t₂ := image₂_subset @@ -559,7 +559,7 @@ theorem singleton_div (a : α) : {a} / s = s.image (a / ·) := theorem singleton_div_singleton (a b : α) : ({a} : Finset α) / {b} = {a / b} := image₂_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem div_subset_div : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ / t₁ ⊆ s₂ / t₂ := image₂_subset @@ -901,6 +901,12 @@ theorem isUnit_coe : IsUnit (s : Set α) ↔ IsUnit s := by @[to_additive (attr := simp)] lemma univ_div_univ [Fintype α] : (univ / univ : Finset α) = univ := by simp [div_eq_mul_inv] +@[to_additive] lemma subset_div_left (ht : 1 ∈ t) : s ⊆ s / t := by + rw [div_eq_mul_inv]; exact subset_mul_left _ <| by simpa + +@[to_additive] lemma inv_subset_div_right (hs : 1 ∈ s) : t⁻¹ ⊆ s / t := by + rw [div_eq_mul_inv]; exact subset_mul_right _ hs + end DivisionMonoid /-- `Finset α` is a commutative division monoid under pointwise operations if `α` is. -/ @@ -1147,7 +1153,7 @@ theorem smul_singleton (b : β) : s • ({b} : Finset β) = s.image (· • b) : theorem singleton_smul_singleton (a : α) (b : β) : ({a} : Finset α) • ({b} : Finset β) = {a • b} := image₂_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image₂_subset @@ -1269,7 +1275,7 @@ theorem singleton_vsub (a : β) : ({a} : Finset β) -ᵥ t = t.image (a -ᵥ ·) theorem singleton_vsub_singleton (a b : β) : ({a} : Finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton -@[mono] +@[mono, gcongr] theorem vsub_subset_vsub : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image₂_subset @@ -1368,7 +1374,7 @@ theorem Nonempty.smul_finset (hs : s.Nonempty) : (a • s).Nonempty := theorem singleton_smul (a : α) : ({a} : Finset α) • t = a • t := image₂_singleton_left -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem smul_finset_subset_smul_finset : s ⊆ t → a • s ⊆ a • t := image_subset_image @@ -1380,6 +1386,10 @@ theorem smul_finset_singleton (b : β) : a • ({b} : Finset β) = {a • b} := 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 _ _ _ diff --git a/Mathlib/Algebra/Group/Pointwise/Set.lean b/Mathlib/Algebra/Group/Pointwise/Set.lean index a1ac52d71f441..1ab1dcd9ea7a1 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set.lean @@ -324,7 +324,7 @@ theorem singleton_mul : {a} * t = (a * ·) '' t := theorem singleton_mul_singleton : ({a} : Set α) * {b} = {a * b} := image2_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem mul_subset_mul : s₁ ⊆ t₁ → s₂ ⊆ t₂ → s₁ * s₂ ⊆ t₁ * t₂ := image2_subset @@ -512,7 +512,7 @@ theorem singleton_div : {a} / t = (· / ·) a '' t := theorem singleton_div_singleton : ({a} : Set α) / {b} = {a / b} := image2_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] theorem div_subset_div : s₁ ⊆ t₁ → s₂ ⊆ t₂ → s₁ / s₂ ⊆ t₁ / t₂ := image2_subset @@ -659,7 +659,7 @@ lemma singleton_smul : ({a} : Set α) • t = a • t := image2_singleton_left @[to_additive (attr := simp high)] lemma singleton_smul_singleton : ({a} : Set α) • ({b} : Set β) = {a • b} := image2_singleton -@[to_additive (attr := mono)] +@[to_additive (attr := mono, gcongr)] lemma smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image2_subset @[to_additive] lemma smul_subset_smul_left : t₁ ⊆ t₂ → s • t₁ ⊆ s • t₂ := image2_subset_left @@ -749,9 +749,7 @@ lemma smul_set_nonempty : (a • s).Nonempty ↔ s.Nonempty := image_nonempty @[to_additive (attr := simp)] lemma smul_set_singleton : a • ({b} : Set β) = {a • b} := image_singleton -@[to_additive] -lemma smul_set_mono : s ⊆ t → a • s ⊆ a • t := - image_subset _ +@[to_additive (attr := gcongr)] lemma smul_set_mono : s ⊆ t → a • s ⊆ a • t := image_subset _ @[to_additive] lemma smul_set_subset_iff : a • s ⊆ t ↔ ∀ ⦃b⦄, b ∈ s → a • b ∈ t := @@ -761,6 +759,10 @@ lemma smul_set_subset_iff : a • s ⊆ t ↔ ∀ ⦃b⦄, b ∈ s → a • b lemma smul_set_union : a • (t₁ ∪ t₂) = a • t₁ ∪ a • t₂ := image_union _ _ _ +@[to_additive] +lemma smul_set_insert (a : α) (b : β) (s : Set β) : a • insert b s = insert (a • b) (a • s) := + image_insert_eq .. + @[to_additive] lemma smul_set_inter_subset : a • (t₁ ∩ t₂) ⊆ a • t₁ ∩ a • t₂ := image_inter_subset _ _ _ @@ -1107,6 +1109,12 @@ theorem isUnit_iff : IsUnit s ↔ ∃ a, s = {a} ∧ IsUnit a := by @[to_additive (attr := simp)] lemma univ_div_univ : (univ / univ : Set α) = univ := by simp [div_eq_mul_inv] +@[to_additive] lemma subset_div_left (ht : 1 ∈ t) : s ⊆ s / t := by + rw [div_eq_mul_inv]; exact subset_mul_left _ <| by simpa + +@[to_additive] lemma inv_subset_div_right (hs : 1 ∈ s) : t⁻¹ ⊆ s / t := by + rw [div_eq_mul_inv]; exact subset_mul_right _ hs + end DivisionMonoid /-- `Set α` is a commutative division monoid under pointwise operations if `α` is. -/ diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 264bf06f72a92..4f9ff4c003ac1 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -607,7 +607,6 @@ theorem tendsto_addHaar_inter_smul_zero_of_density_zero_aux1 (s : Set E) (x : E) rintro r (rpos : 0 < r) rw [← affinity_unitClosedBall rpos.le, singleton_add, ← image_vadd] gcongr - exact smul_set_mono t_bound have B : Tendsto (fun r : ℝ => μ (closedBall x r) / μ ({x} + r • u)) (𝓝[>] 0) (𝓝 (μ (closedBall x 1) / μ ({x} + u))) := by From 874e7093b54fab979258ff2e675dffd11c169b47 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Sep 2024 21:49:03 +0000 Subject: [PATCH 023/174] chore(Pointwise): Rename `Set` to `Set.Basic` (#17234) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Move: * `Algebra.Group.Pointwise.Set` → `Algebra.Group.Pointwise.Set.Basic` * `Algebra.Group.Pointwise.Finset.NatCard` → `Algebra.Group.Pointwise.Set.Card` * `Algebra.GroupWithZero.Pointwise.Set` → `Algebra.GroupWithZero.Pointwise.Set.Basic` This is making space for a new file `Algebra.GroupWithZero.Pointwise.Set.Card` in #17231. `Algebra.Group.Pointwise.Finset` already has this `.Basic` naming scheme. --- Mathlib.lean | 6 +++--- Mathlib/Algebra/AddTorsor.lean | 2 +- .../Algebra/Group/Pointwise/{Set.lean => Set/Basic.lean} | 0 .../Group/Pointwise/{Finset/NatCard.lean => Set/Card.lean} | 0 .../GroupWithZero/Pointwise/{Set.lean => Set/Basic.lean} | 2 +- Mathlib/Algebra/Order/Field/Pointwise.lean | 2 +- Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean | 2 +- Mathlib/Algebra/Order/Interval/Basic.lean | 2 +- Mathlib/Algebra/Order/Module/Pointwise.lean | 2 +- Mathlib/Algebra/Ring/Pointwise/Set.lean | 2 +- Mathlib/Algebra/Star/Conjneg.lean | 2 +- Mathlib/Algebra/Star/Pointwise.lean | 2 +- Mathlib/Combinatorics/Additive/Dissociation.lean | 2 +- Mathlib/Combinatorics/Additive/FreimanHom.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Circulant.lean | 2 +- Mathlib/Data/Finset/MulAntidiagonal.lean | 2 +- Mathlib/Data/Finset/SMulAntidiagonal.lean | 2 +- Mathlib/Data/Set/Pointwise/BigOperators.lean | 2 +- Mathlib/Data/Set/Pointwise/BoundedMul.lean | 2 +- Mathlib/Data/Set/Pointwise/Finite.lean | 2 +- Mathlib/Data/Set/Pointwise/Interval.lean | 2 +- Mathlib/Data/Set/Pointwise/ListOfFn.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 2 +- Mathlib/Data/Set/Semiring.lean | 2 +- Mathlib/GroupTheory/GroupAction/Pointwise.lean | 2 +- Mathlib/GroupTheory/GroupAction/Support.lean | 2 +- 26 files changed, 26 insertions(+), 26 deletions(-) rename Mathlib/Algebra/Group/Pointwise/{Set.lean => Set/Basic.lean} (100%) rename Mathlib/Algebra/Group/Pointwise/{Finset/NatCard.lean => Set/Card.lean} (100%) rename Mathlib/Algebra/GroupWithZero/Pointwise/{Set.lean => Set/Basic.lean} (97%) diff --git a/Mathlib.lean b/Mathlib.lean index 6894a6f609343..9a8dfc434a318 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -267,8 +267,8 @@ import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pi.Lemmas import Mathlib.Algebra.Group.Pointwise.Finset.Basic import Mathlib.Algebra.Group.Pointwise.Finset.Interval -import Mathlib.Algebra.Group.Pointwise.Finset.NatCard -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic +import Mathlib.Algebra.Group.Pointwise.Set.Card import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Semiconj.Basic import Mathlib.Algebra.Group.Semiconj.Defs @@ -322,7 +322,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.Set +import Mathlib.Algebra.GroupWithZero.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Prod import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.GroupWithZero.ULift diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index f9baa6191726f..9d3b4b23f9e67 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Basic -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic /-! # Torsors of additive group actions diff --git a/Mathlib/Algebra/Group/Pointwise/Set.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean similarity index 100% rename from Mathlib/Algebra/Group/Pointwise/Set.lean rename to Mathlib/Algebra/Group/Pointwise/Set/Basic.lean diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/NatCard.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean similarity index 100% rename from Mathlib/Algebra/Group/Pointwise/Finset/NatCard.lean rename to Mathlib/Algebra/Group/Pointwise/Set/Card.lean diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Set.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Basic.lean similarity index 97% rename from Mathlib/Algebra/GroupWithZero/Pointwise/Set.lean rename to Mathlib/Algebra/GroupWithZero/Pointwise/Set/Basic.lean index 653b0ea020618..ec178273360b2 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Set.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Basic.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, Floris van Doorn -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Basic /-! diff --git a/Mathlib/Algebra/Order/Field/Pointwise.lean b/Mathlib/Algebra/Order/Field/Pointwise.lean index 8f2b1867f0d73..0b5b49e3c6bb9 100644 --- a/Mathlib/Algebra/Order/Field/Pointwise.lean +++ b/Mathlib/Algebra/Order/Field/Pointwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Yaël Dillies -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.SMulWithZero diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index e6a5407931b80..1b1e0d3b70882 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.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 -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Algebra.Order.Monoid.Unbundled.OrderDual import Mathlib.Order.Bounds.OrderIso diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index 53c70afe40a7b..dcd8a5d261e09 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.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.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Order.Interval.Basic diff --git a/Mathlib/Algebra/Order/Module/Pointwise.lean b/Mathlib/Algebra/Order/Module/Pointwise.lean index a8bede48a1df5..147291c9cdcef 100644 --- a/Mathlib/Algebra/Order/Module/Pointwise.lean +++ b/Mathlib/Algebra/Order/Module/Pointwise.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.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Module.Defs import Mathlib.Order.Bounds.OrderIso diff --git a/Mathlib/Algebra/Ring/Pointwise/Set.lean b/Mathlib/Algebra/Ring/Pointwise/Set.lean index 4da48ed12ad44..8c7e70fe1622b 100644 --- a/Mathlib/Algebra/Ring/Pointwise/Set.lean +++ b/Mathlib/Algebra/Ring/Pointwise/Set.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, Floris van Doorn -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Ring.Defs /-! diff --git a/Mathlib/Algebra/Star/Conjneg.lean b/Mathlib/Algebra/Star/Conjneg.lean index 1fb96695c6371..9b9c3b54defe4 100644 --- a/Mathlib/Algebra/Star/Conjneg.lean +++ b/Mathlib/Algebra/Star/Conjneg.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Pi -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Star.Pi /-! diff --git a/Mathlib/Algebra/Star/Pointwise.lean b/Mathlib/Algebra/Star/Pointwise.lean index 58c6df130b0b7..6c37067eee6ab 100644 --- a/Mathlib/Algebra/Star/Pointwise.lean +++ b/Mathlib/Algebra/Star/Pointwise.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.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Star.Basic import Mathlib.Data.Set.Finite diff --git a/Mathlib/Combinatorics/Additive/Dissociation.lean b/Mathlib/Combinatorics/Additive/Dissociation.lean index d7dfe231cee32..f22b05d75f55d 100644 --- a/Mathlib/Combinatorics/Additive/Dissociation.lean +++ b/Mathlib/Combinatorics/Additive/Dissociation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Data.Fintype.Card diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index 6b1ff91dbf629..dee5a50bf940d 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Order.BigOperators.Group.Multiset import Mathlib.Data.ZMod.Defs diff --git a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean index dabe8553d196a..121f720ec7a12 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Iván Renison, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Iván Renison, Bhavik Mehta -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Combinatorics.SimpleGraph.Hasse /-! diff --git a/Mathlib/Data/Finset/MulAntidiagonal.lean b/Mathlib/Data/Finset/MulAntidiagonal.lean index d361b2e348cd7..534b98c618c6b 100644 --- a/Mathlib/Data/Finset/MulAntidiagonal.lean +++ b/Mathlib/Data/Finset/MulAntidiagonal.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yaël Dillies -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Data.Set.MulAntidiagonal diff --git a/Mathlib/Data/Finset/SMulAntidiagonal.lean b/Mathlib/Data/Finset/SMulAntidiagonal.lean index 20da3b77568f0..53d0c959253ec 100644 --- a/Mathlib/Data/Finset/SMulAntidiagonal.lean +++ b/Mathlib/Data/Finset/SMulAntidiagonal.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Scott Carnahan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Data.Set.SMulAntidiagonal /-! diff --git a/Mathlib/Data/Set/Pointwise/BigOperators.lean b/Mathlib/Data/Set/Pointwise/BigOperators.lean index 2cc66bb66f328..7772e8e35548f 100644 --- a/Mathlib/Data/Set/Pointwise/BigOperators.lean +++ b/Mathlib/Data/Set/Pointwise/BigOperators.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic /-! # Results about pointwise operations on sets and big operators. diff --git a/Mathlib/Data/Set/Pointwise/BoundedMul.lean b/Mathlib/Data/Set/Pointwise/BoundedMul.lean index 0472b092e5f72..2d630b685665e 100644 --- a/Mathlib/Data/Set/Pointwise/BoundedMul.lean +++ b/Mathlib/Data/Set/Pointwise/BoundedMul.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 KudryashovJ -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Monoid.Defs /-! diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Data/Set/Pointwise/Finite.lean index dc7f3efc64c90..934fd477ecac1 100644 --- a/Mathlib/Data/Set/Pointwise/Finite.lean +++ b/Mathlib/Data/Set/Pointwise/Finite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ import Mathlib.Algebra.Group.Action.Basic -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Data.Set.Finite /-! # Finiteness lemmas for pointwise operations on sets -/ diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 291fcb810c4c8..1b83b2cd5269f 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.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, Patrick Massot -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Group.MinMax import Mathlib.Algebra.Order.Interval.Set.Monoid diff --git a/Mathlib/Data/Set/Pointwise/ListOfFn.lean b/Mathlib/Data/Set/Pointwise/ListOfFn.lean index 309b0b70bb701..21e7712a88181 100644 --- a/Mathlib/Data/Set/Pointwise/ListOfFn.lean +++ b/Mathlib/Data/Set/Pointwise/ListOfFn.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.BigOperators.Group.List -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Data.List.OfFn /-! diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 7e53675dbbac8..e3ed79c944b46 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ import Mathlib.Algebra.Group.Pi.Basic -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs import Mathlib.Data.Set.Pairwise.Basic diff --git a/Mathlib/Data/Set/Semiring.lean b/Mathlib/Data/Set/Semiring.lean index 65db6ecc8dee6..e6ff5cf305201 100644 --- a/Mathlib/Data/Set/Semiring.lean +++ b/Mathlib/Data/Set/Semiring.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Order.Kleene import Mathlib.Algebra.Order.Ring.Canonical diff --git a/Mathlib/GroupTheory/GroupAction/Pointwise.lean b/Mathlib/GroupTheory/GroupAction/Pointwise.lean index b5537fe937116..4a0bf768bba53 100644 --- a/Mathlib/GroupTheory/GroupAction/Pointwise.lean +++ b/Mathlib/GroupTheory/GroupAction/Pointwise.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, Antoine Chambert-Loir -/ -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.GroupTheory.GroupAction.Hom /-! diff --git a/Mathlib/GroupTheory/GroupAction/Support.lean b/Mathlib/GroupTheory/GroupAction/Support.lean index c2bd6a25eddf5..392881eb03ad9 100644 --- a/Mathlib/GroupTheory/GroupAction/Support.lean +++ b/Mathlib/GroupTheory/GroupAction/Support.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.Action.Basic -import Mathlib.Algebra.Group.Pointwise.Set +import Mathlib.Algebra.Group.Pointwise.Set.Basic /-! # Support of an element under an action action From d453819bc00b126955d12d64d4952e8eb1e21993 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 28 Sep 2024 23:34:14 +0000 Subject: [PATCH 024/174] feat: the real exponential is analytic (#17218) Currently, mathlib proves that the complex exponential is analytic as a consequence of its differentiability, quite late in the import hierarchy -- and this is specific to complexes. On the other hand, we also know that the exponential on a general normed algebra is analytic, thanks to its power series expansion. We switch to deduce the former from the latter, earlier in the import hierarchy, and use this occasion to also prove analyticity of the real exponential in the same file. --- Mathlib/Analysis/Analytic/Constructions.lean | 60 ++++++++++++ .../Analysis/Calculus/FDeriv/Analytic.lean | 6 ++ .../SpecialFunctions/Complex/Analytic.lean | 27 +----- .../Analysis/SpecialFunctions/ExpDeriv.lean | 91 +++++++++++++++---- 4 files changed, 142 insertions(+), 42 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 7ef5b6a75773b..5b8b4b639daa3 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -649,6 +649,61 @@ lemma AnalyticOnNhd.pow {f : E → A} {s : Set E} (hf : AnalyticOnNhd 𝕜 f s) AnalyticOnNhd 𝕜 (fun x ↦ f x ^ n) s := fun _ m ↦ (hf _ m).pow n + +/-! +### Restriction of scalars +-/ + +section + +variable {𝕜' : Type*} [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] + [NormedSpace 𝕜' E] [IsScalarTower 𝕜 𝕜' E] + [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] + {f : E → F} {p : FormalMultilinearSeries 𝕜' E F} {x : E} {s : Set E} {r : ℝ≥0∞} + +lemma HasFPowerSeriesWithinOnBall.restrictScalars (hf : HasFPowerSeriesWithinOnBall f p s x r) : + HasFPowerSeriesWithinOnBall f (p.restrictScalars 𝕜) s x r := + ⟨hf.r_le.trans (FormalMultilinearSeries.radius_le_of_le (fun n ↦ by simp)), hf.r_pos, hf.hasSum⟩ + +lemma HasFPowerSeriesOnBall.restrictScalars (hf : HasFPowerSeriesOnBall f p x r) : + HasFPowerSeriesOnBall f (p.restrictScalars 𝕜) x r := + ⟨hf.r_le.trans (FormalMultilinearSeries.radius_le_of_le (fun n ↦ by simp)), hf.r_pos, hf.hasSum⟩ + +lemma HasFPowerSeriesWithinAt.restrictScalars (hf : HasFPowerSeriesWithinAt f p s x) : + HasFPowerSeriesWithinAt f (p.restrictScalars 𝕜) s x := by + rcases hf with ⟨r, hr⟩ + exact ⟨r, hr.restrictScalars⟩ + +lemma HasFPowerSeriesAt.restrictScalars (hf : HasFPowerSeriesAt f p x) : + HasFPowerSeriesAt f (p.restrictScalars 𝕜) x := by + rcases hf with ⟨r, hr⟩ + exact ⟨r, hr.restrictScalars⟩ + +lemma AnalyticWithinAt.restrictScalars (hf : AnalyticWithinAt 𝕜' f s x) : + AnalyticWithinAt 𝕜 f s x := by + rcases hf with ⟨p, hp⟩ + exact ⟨p.restrictScalars 𝕜, hp.restrictScalars⟩ + +lemma AnalyticAt.restrictScalars (hf : AnalyticAt 𝕜' f x) : + AnalyticAt 𝕜 f x := by + rcases hf with ⟨p, hp⟩ + exact ⟨p.restrictScalars 𝕜, hp.restrictScalars⟩ + +lemma AnalyticOn.restrictScalars (hf : AnalyticOn 𝕜' f s) : + AnalyticOn 𝕜 f s := + fun x hx ↦ (hf x hx).restrictScalars + +lemma AnalyticOnNhd.restrictScalars (hf : AnalyticOnNhd 𝕜' f s) : + AnalyticOnNhd 𝕜 f s := + fun x hx ↦ (hf x hx).restrictScalars + +end + + +/-! +### Inversion is analytic +-/ + section Geometric variable (𝕜 A : Type*) [NontriviallyNormedField 𝕜] [NormedRing A] [NormedAlgebra 𝕜 A] @@ -757,6 +812,11 @@ lemma analyticAt_inverse {𝕜 : Type*} [NontriviallyNormedField 𝕜] exact analyticAt_inverse_one_sub 𝕜 A · exact analyticAt_const.sub (analyticAt_const.mul analyticAt_id) +lemma analyticOnNhd_inverse {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] [HasSummableGeomSeries A] : + AnalyticOnNhd 𝕜 Ring.inverse {x : A | IsUnit x} := + fun _ hx ↦ analyticAt_inverse (IsUnit.unit hx) + lemma hasFPowerSeriesOnBall_inv_one_sub (𝕜 𝕝 : Type*) [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕝] [NormedAlgebra 𝕜 𝕝] : HasFPowerSeriesOnBall (fun x : 𝕝 ↦ (1 - x)⁻¹) (formalMultilinearSeries_geometric 𝕜 𝕝) 0 1 := by diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 896125484853c..92be2e3cc4f08 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -138,6 +138,12 @@ theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (fun m _ ↦ (H.iteratedFDeriv m).differentiableOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) +/-- An analytic function on the whole space is infinitely differentiable there. -/ +theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : ℕ∞} : + ContDiff 𝕜 n f := by + rw [← contDiffOn_univ] + exact h.contDiffOn + theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : ℕ∞} : ContDiffAt 𝕜 n f x := by obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean index dfa5088bf9e96..d717114a27b82 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean @@ -11,7 +11,7 @@ import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv /-! # Various complex special functions are analytic -`exp`, `log`, and `cpow` are analytic, since they are differentiable. +`log`, and `cpow` are analytic, since they are differentiable. -/ open Complex Set @@ -20,31 +20,6 @@ open scoped Topology variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] variable {f g : E → ℂ} {z : ℂ} {x : E} {s : Set E} -/-- `exp` is entire -/ -theorem analyticOnNhd_cexp : AnalyticOnNhd ℂ exp univ := by - rw [analyticOnNhd_univ_iff_differentiable]; exact differentiable_exp - -theorem analyticOn_cexp : AnalyticOn ℂ exp univ := analyticOnNhd_cexp.analyticOn - -/-- `exp` is analytic at any point -/ -theorem analyticAt_cexp : AnalyticAt ℂ exp z := - analyticOnNhd_cexp z (mem_univ _) - -/-- `exp ∘ f` is analytic -/ -theorem AnalyticAt.cexp (fa : AnalyticAt ℂ f x) : AnalyticAt ℂ (fun z ↦ exp (f z)) x := - analyticAt_cexp.comp fa - -theorem AnalyticWithinAt.cexp (fa : AnalyticWithinAt ℂ f s x) : - AnalyticWithinAt ℂ (fun z ↦ exp (f z)) s x := - analyticAt_cexp.comp_analyticWithinAt fa - -/-- `exp ∘ f` is analytic -/ -theorem AnalyticOnNhd.cexp (fs : AnalyticOnNhd ℂ f s) : AnalyticOnNhd ℂ (fun z ↦ exp (f z)) s := - fun z n ↦ analyticAt_cexp.comp (fs z n) - -theorem AnalyticOn.cexp (fs : AnalyticOn ℂ f s) : AnalyticOn ℂ (fun z ↦ exp (f z)) s := - analyticOnNhd_cexp.comp_analyticOn fs (mapsTo_univ _ _) - /-- `log` is analytic away from nonpositive reals -/ theorem analyticAt_clog (m : z ∈ slitPlane) : AnalyticAt ℂ log z := by rw [analyticAt_iff_eventually_differentiableAt] diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 31ff5558fe80a..8506d9e055ab1 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne import Mathlib.Analysis.Complex.RealDeriv import Mathlib.Analysis.Calculus.ContDiff.RCLike import Mathlib.Analysis.Calculus.IteratedDeriv.Lemmas +import Mathlib.Analysis.SpecialFunctions.Exponential /-! # Complex and real exponential @@ -24,6 +25,41 @@ open scoped Topology /-! ## `Complex.exp` -/ +section + +open Complex + +variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] +variable {f g : E → ℂ} {z : ℂ} {x : E} {s : Set E} + +/-- `exp` is entire -/ +theorem analyticOnNhd_cexp : AnalyticOnNhd ℂ exp univ := by + rw [Complex.exp_eq_exp_ℂ] + exact fun x _ ↦ NormedSpace.exp_analytic x + +theorem analyticOn_cexp : AnalyticOn ℂ exp univ := analyticOnNhd_cexp.analyticOn + +/-- `exp` is analytic at any point -/ +theorem analyticAt_cexp : AnalyticAt ℂ exp z := + analyticOnNhd_cexp z (mem_univ _) + +/-- `exp ∘ f` is analytic -/ +theorem AnalyticAt.cexp (fa : AnalyticAt ℂ f x) : AnalyticAt ℂ (fun z ↦ exp (f z)) x := + analyticAt_cexp.comp fa + +theorem AnalyticWithinAt.cexp (fa : AnalyticWithinAt ℂ f s x) : + AnalyticWithinAt ℂ (fun z ↦ exp (f z)) s x := + analyticAt_cexp.comp_analyticWithinAt fa + +/-- `exp ∘ f` is analytic -/ +theorem AnalyticOnNhd.cexp (fs : AnalyticOnNhd ℂ f s) : AnalyticOnNhd ℂ (fun z ↦ exp (f z)) s := + fun z n ↦ analyticAt_cexp.comp (fs z n) + +theorem AnalyticOn.cexp (fs : AnalyticOn ℂ f s) : AnalyticOn ℂ (fun z ↦ exp (f z)) s := + analyticOnNhd_cexp.comp_analyticOn fs (mapsTo_univ _ _) + +end + namespace Complex variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜 ℂ] @@ -52,17 +88,8 @@ theorem iter_deriv_exp : ∀ n : ℕ, deriv^[n] exp = exp | 0 => rfl | n + 1 => by rw [iterate_succ_apply, deriv_exp, iter_deriv_exp n] -theorem contDiff_exp : ∀ {n}, ContDiff 𝕜 n exp := by - -- Porting note: added `@` due to `∀ {n}` weirdness above - refine @(contDiff_all_iff_nat.2 fun n => ?_) - have : ContDiff ℂ (↑n) exp := by - induction n with - | zero => exact contDiff_zero.2 continuous_exp - | succ n ihn => - rw [contDiff_succ_iff_deriv] - use differentiable_exp - rwa [deriv_exp] - exact this.restrict_scalars 𝕜 +theorem contDiff_exp {n : ℕ∞} : ContDiff 𝕜 n exp := + analyticOnNhd_cexp.restrictScalars.contDiff theorem hasStrictDerivAt_exp (x : ℂ) : HasStrictDerivAt exp (exp x) x := contDiff_exp.contDiffAt.hasStrictDerivAt' (hasDerivAt_exp x) le_rfl @@ -156,12 +183,44 @@ theorem iteratedDeriv_cexp_const_mul (n : ℕ) (c : ℂ) : (iteratedDeriv n fun s : ℂ => exp (c * s)) = fun s => c ^ n * exp (c * s) := by rw [iteratedDeriv_const_mul contDiff_exp, iteratedDeriv_eq_iterate, iter_deriv_exp] - /-! ## `Real.exp` -/ -namespace Real +section + +open Real + +variable {x : ℝ} {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {f : E → ℝ} {s : Set E} + +/-- `exp` is entire -/ +theorem analyticOnNhd_rexp : AnalyticOnNhd ℝ exp univ := by + rw [Real.exp_eq_exp_ℝ] + exact fun x _ ↦ NormedSpace.exp_analytic x + +theorem analyticOn_rexp : AnalyticOn ℝ exp univ := analyticOnNhd_rexp.analyticOn + +/-- `exp` is analytic at any point -/ +theorem analyticAt_rexp : AnalyticAt ℝ exp x := + analyticOnNhd_rexp x (mem_univ _) -variable {x y z : ℝ} +/-- `exp ∘ f` is analytic -/ +theorem AnalyticAt.rexp {x : E} (fa : AnalyticAt ℝ f x) : AnalyticAt ℝ (fun z ↦ exp (f z)) x := + analyticAt_rexp.comp fa + +theorem AnalyticWithinAt.rexp {x : E} (fa : AnalyticWithinAt ℝ f s x) : + AnalyticWithinAt ℝ (fun z ↦ exp (f z)) s x := + analyticAt_rexp.comp_analyticWithinAt fa + +/-- `exp ∘ f` is analytic -/ +theorem AnalyticOnNhd.rexp {s : Set E} (fs : AnalyticOnNhd ℝ f s) : + AnalyticOnNhd ℝ (fun z ↦ exp (f z)) s := + fun z n ↦ analyticAt_rexp.comp (fs z n) + +theorem AnalyticOn.rexp (fs : AnalyticOn ℝ f s) : AnalyticOn ℝ (fun z ↦ exp (f z)) s := + analyticOnNhd_rexp.comp_analyticOn fs (mapsTo_univ _ _) + +end + +namespace Real theorem hasStrictDerivAt_exp (x : ℝ) : HasStrictDerivAt exp (exp x) x := (Complex.hasStrictDerivAt_exp x).real_of_complex @@ -169,12 +228,12 @@ theorem hasStrictDerivAt_exp (x : ℝ) : HasStrictDerivAt exp (exp x) x := theorem hasDerivAt_exp (x : ℝ) : HasDerivAt exp (exp x) x := (Complex.hasDerivAt_exp x).real_of_complex -theorem contDiff_exp {n} : ContDiff ℝ n exp := +theorem contDiff_exp {n : ℕ∞} : ContDiff ℝ n exp := Complex.contDiff_exp.real_of_complex theorem differentiable_exp : Differentiable ℝ exp := fun x => (hasDerivAt_exp x).differentiableAt -theorem differentiableAt_exp : DifferentiableAt ℝ exp x := +theorem differentiableAt_exp {x : ℝ} : DifferentiableAt ℝ exp x := differentiable_exp x @[simp] From 2845f11d96267fc699a8c1024e3caf998d4b3fdc Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 28 Sep 2024 23:34:15 +0000 Subject: [PATCH 025/174] feat(SeparationQuotient): add missing instances (#17239) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also prove `SeparationQuotient.postcomp_mkCLM_surjective`. I'm going to need these instances and lemmas to generalize `CompleteSpace (E →L[K] F)` to topological vector spaces without assuming that `F` is a Hausdorff space, see #17244. --- .../Topology/Algebra/SeparationQuotient.lean | 28 ++++++++++++++++++- Mathlib/Topology/Inseparable.lean | 3 ++ 2 files changed, 30 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient.lean index 140e089dc9563..19e68ba576e4f 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient.lean @@ -3,8 +3,9 @@ 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.Algebra.Module.Basic import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Maps.OpenQuotient /-! # Algebraic operations on `SeparationQuotient` @@ -63,6 +64,12 @@ instance instIsScalarTower [SMul M N] [ContinuousConstSMul N X] [IsScalarTower M 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] + exact continuous_mk.comp continuous_smul + instance instSMulZeroClass {M X : Type*} [Zero X] [SMulZeroClass M X] [TopologicalSpace X] [ContinuousConstSMul M X] : SMulZeroClass M (SeparationQuotient X) := ZeroHom.smulZeroClass ⟨mk, mk_zero⟩ mk_smul @@ -190,6 +197,17 @@ instance instCommGroup [CommGroup G] [TopologicalGroup G] : CommGroup (Separatio end Group +section UniformGroup + +@[to_additive] +instance instUniformGroup {G : Type*} [Group G] [UniformSpace G] [UniformGroup G] : + UniformGroup (SeparationQuotient G) where + uniformContinuous_div := by + rw [uniformContinuous_dom₂] + exact uniformContinuous_mk.comp uniformContinuous_div + +end UniformGroup + section MonoidWithZero variable {M₀ : Type*} [TopologicalSpace M₀] @@ -391,6 +409,14 @@ theorem mk_outCLM (x : SeparationQuotient E) : mk (outCLM K E x) = 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 _) diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 392ac4f52c773..f122cb8067af0 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -548,6 +548,9 @@ theorem preimage_image_mk_open (hs : IsOpen s) : mk ⁻¹' (mk '' s) = s := by theorem isOpenMap_mk : IsOpenMap (mk : X → SeparationQuotient X) := fun s hs => quotientMap_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⟩ + theorem preimage_image_mk_closed (hs : IsClosed s) : mk ⁻¹' (mk '' s) = s := by refine Subset.antisymm ?_ (subset_preimage_image _ _) rintro x ⟨y, hys, hxy⟩ From ba0116042e05041f0b9efa7ec651212efb69529b Mon Sep 17 00:00:00 2001 From: Colin Jones <137637335+Colin166@users.noreply.github.com> Date: Sun, 29 Sep 2024 00:04:19 +0000 Subject: [PATCH 026/174] feat(NumberTheory/FactorisationProperties): abundant, pseudoperfect, deficient, and weird numbers (#14269) --- Mathlib.lean | 1 + .../NumberTheory/FactorisationProperties.lean | 198 ++++++++++++++++++ docs/references.bib | 11 + 3 files changed, 210 insertions(+) create mode 100644 Mathlib/NumberTheory/FactorisationProperties.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9a8dfc434a318..1acceb28d4138 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3422,6 +3422,7 @@ import Mathlib.NumberTheory.EulerProduct.DirichletLSeries import Mathlib.NumberTheory.FLT.Basic import Mathlib.NumberTheory.FLT.Four import Mathlib.NumberTheory.FLT.Three +import Mathlib.NumberTheory.FactorisationProperties import Mathlib.NumberTheory.FermatPsp import Mathlib.NumberTheory.FrobeniusNumber import Mathlib.NumberTheory.FunctionField diff --git a/Mathlib/NumberTheory/FactorisationProperties.lean b/Mathlib/NumberTheory/FactorisationProperties.lean new file mode 100644 index 0000000000000..28cabcc7fd958 --- /dev/null +++ b/Mathlib/NumberTheory/FactorisationProperties.lean @@ -0,0 +1,198 @@ +/- +Copyright (c) 2024 Colin Jones. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Colin Jones +-/ +import Mathlib.Algebra.GeomSum +import Mathlib.Algebra.IsPrimePow +import Mathlib.NumberTheory.Divisors +import Mathlib.Tactic.FinCases +import Mathlib.Tactic.NormNum.Prime + +/-! +# Factorisation properties of natural numbers + +This file defines abundant, pseudoperfect, deficient, and weird numbers and formalizes their +relations with prime and perfect numbers. + +## Main Definitions + +* `Nat.Abundant`: a natural number `n` is _abundant_ if the sum of its proper divisors is greater + than `n` +* `Nat.Pseudoperfect`: a natural number `n` is _pseudoperfect_ if the sum of a subset of its proper + divisors equals `n` +* `Nat.Deficient`: a natural number `n` is _deficient_ if the sum of its proper divisors is less + than `n` +* `Nat.Weird`: a natural number is _weird_ if it is abundant but not pseudoperfect + +## Main Results + +* `Nat.deficient_or_perfect_or_abundant`: A positive natural number is either deficient, + perfect, or abundant. +* `Nat.Prime.deficient`: All prime natural numbers are deficient. +* `Nat.infinite_deficient`: There are infinitely many deficient numbers. +* `Nat.Prime.deficient_pow`: Any natural number power of a prime is deficient. + +## Implementation Notes +* Zero is not included in any of the definitions and these definitions only apply to natural + numbers greater than zero. + +## References +* [R. W. Prielipp, *PERFECT NUMBERS, ABUNDANT NUMBERS, AND DEFICIENT NUMBERS*][Prielipp1970] + +## Tags + +abundant, deficient, weird, pseudoperfect +-/ + +open Finset + +namespace Nat + +variable {n m p : ℕ} + +/-- `n : ℕ` is _abundant_ if the sum of the proper divisors of `n` is greater than `n`. -/ +def Abundant (n : ℕ) : Prop := n < ∑ i ∈ properDivisors n, i + +/-- `n : ℕ` is _deficient_ if the sum of the proper divisors of `n` is less than `n`. -/ +def Deficient (n : ℕ) : Prop := ∑ i ∈ properDivisors n, i < n + +/-- A positive natural number `n` is _pseudoperfect_ if there exists a subset of the proper + divisors of `n` such that the sum of that subset is equal to `n`. -/ +def Pseudoperfect (n : ℕ) : Prop := + 0 < n ∧ ∃ s ⊆ properDivisors n, ∑ i ∈ s, i = n + +/-- `n : ℕ` is a _weird_ number if and only if it is abundant but not pseudoperfect. -/ +def Weird (n : ℕ) : Prop := Abundant n ∧ ¬ Pseudoperfect n + +theorem not_pseudoperfect_iff_forall : + ¬ Pseudoperfect n ↔ n = 0 ∨ ∀ s ⊆ properDivisors n, ∑ i ∈ s, i ≠ n := by + rw [Pseudoperfect, not_and_or] + simp only [not_lt, nonpos_iff_eq_zero, mem_powerset, not_exists, not_and, ne_eq] + +theorem deficient_one : Deficient 1 := zero_lt_one +theorem deficient_two : Deficient 2 := one_lt_two +theorem deficient_three : Deficient 3 := by norm_num [Deficient] + +theorem abundant_twelve : Abundant 12 := by + rw [Abundant, show properDivisors 12 = {1,2,3,4,6} by rfl] + norm_num + +set_option maxRecDepth 1730 in +theorem weird_seventy : Weird 70 := by + rw [Weird, Abundant, not_pseudoperfect_iff_forall] + have h : properDivisors 70 = {1, 2, 5, 7, 10, 14, 35} := by rfl + constructor + · rw [h] + repeat norm_num + · rw [h] + right + intro s hs + have hs' := mem_powerset.mpr hs + fin_cases hs' <;> decide + +lemma deficient_iff_not_abundant_and_not_perfect (hn : n ≠ 0) : + Deficient n ↔ ¬ Abundant n ∧ ¬ Perfect n := by + dsimp only [Perfect, Abundant, Deficient] + omega + +lemma perfect_iff_not_abundant_and_not_deficient (hn : 0 ≠ n) : + Perfect n ↔ ¬ Abundant n ∧ ¬ Deficient n := by + dsimp only [Perfect, Abundant, Deficient] + omega + +lemma abundant_iff_not_perfect_and_not_deficient (hn : 0 ≠ n) : + Abundant n ↔ ¬ Perfect n ∧ ¬ Deficient n := by + dsimp only [Perfect, Abundant, Deficient] + omega + +/-- A positive natural number is either deficient, perfect, or abundant -/ +theorem deficient_or_perfect_or_abundant (hn : 0 ≠ n) : + Deficient n ∨ Abundant n ∨ Perfect n := by + dsimp only [Perfect, Abundant, Deficient] + omega + +theorem Perfect.pseudoperfect (h : Perfect n) : Pseudoperfect n := + ⟨h.2, ⟨properDivisors n, ⟨fun ⦃_⦄ a ↦ a, h.1⟩⟩⟩ + +theorem Prime.not_abundant (h : Prime n) : ¬ Abundant n := + fun h1 ↦ (h.one_lt.trans h1).ne' (sum_properDivisors_eq_one_iff_prime.mpr h) + +theorem Prime.not_weird (h : Prime n) : ¬ Weird n := by + simp only [Nat.Weird, not_and_or] + left + exact h.not_abundant + +theorem Prime.not_pseudoperfect (h : Prime p) : ¬ Pseudoperfect p := by + simp_rw [not_pseudoperfect_iff_forall, ← mem_powerset, + show p.properDivisors.powerset = {∅, {1}} by rw [Prime.properDivisors h]; rfl] + refine Or.inr (fun s hs ↦ ?_) + fin_cases hs <;> + simp only [sum_empty, sum_singleton] <;> + linarith [Prime.one_lt h] + +theorem Prime.not_perfect (h : Prime p) : ¬ Perfect p := by + have h1 := Prime.not_pseudoperfect h + revert h1 + exact not_imp_not.mpr (Perfect.pseudoperfect) + +/-- Any natural number power of a prime is deficient -/ +theorem Prime.deficient_pow (h : Prime n) : Deficient (n ^ m) := by + rcases Nat.eq_zero_or_pos m with (rfl | _) + · simpa using deficient_one + · have h1 : (n ^ m).properDivisors = image (n ^ ·) (range m) := by + apply subset_antisymm <;> intro a + · simp only [mem_properDivisors, mem_image, mem_range, dvd_prime_pow h] + rintro ⟨⟨t, ht, rfl⟩, ha'⟩ + exact ⟨t, lt_of_le_of_ne ht (fun ht' ↦ lt_irrefl _ (ht' ▸ ha')), rfl⟩ + · simp only [mem_image, mem_range, mem_properDivisors, forall_exists_index, and_imp] + intro x hx hy + constructor + · rw [← hy, dvd_prime_pow h] + exact ⟨x, Nat.le_of_succ_le hx, rfl⟩ + · rw [← hy] + exact (Nat.pow_lt_pow_iff_right (Prime.two_le h)).mpr hx + have h2 : ∑ i in image (fun x => n ^ x) (range m), i = ∑ i in range m, n^i := by + rw [Finset.sum_image] + rintro x _ y _ + apply pow_injective_of_not_isUnit h.not_unit <| Prime.ne_zero h + rw [Deficient, h1, h2] + calc + ∑ i ∈ range m, n ^ i + = (n ^ m - 1) / (n - 1) := (Nat.geomSum_eq (Prime.two_le h) _) + _ ≤ (n ^ m - 1) := Nat.div_le_self (n ^ m - 1) (n - 1) + _ < n ^ m := sub_lt (pow_pos (Prime.pos h) m) (Nat.one_pos) + +theorem _root_.IsPrimePow.deficient (h : IsPrimePow n) : Deficient n := by + obtain ⟨p, k, hp, -, rfl⟩ := h + exact hp.nat_prime.deficient_pow + +theorem Prime.deficient (h : Prime n) : Deficient n := by + rw [← pow_one n] + exact h.deficient_pow + +/-- There exists infinitely many deficient numbers -/ +theorem infinite_deficient : {n : ℕ | n.Deficient}.Infinite := by + rw [Set.infinite_iff_exists_gt] + intro a + obtain ⟨b, h1, h2⟩ := exists_infinite_primes a.succ + exact ⟨b, h2.deficient, h1⟩ + +theorem infinite_even_deficient : {n : ℕ | Even n ∧ n.Deficient}.Infinite := by + rw [Set.infinite_iff_exists_gt] + intro n + use 2 ^ (n + 1) + constructor + · exact ⟨⟨2 ^ n, by ring⟩, prime_two.deficient_pow⟩ + · calc + n ≤ 2 ^ n := Nat.le_of_lt (lt_two_pow n) + _ < 2 ^ (n + 1) := (Nat.pow_lt_pow_iff_right (Nat.one_lt_two)).mpr (lt_add_one n) + +theorem infinite_odd_deficient : {n : ℕ | Odd n ∧ n.Deficient}.Infinite := by + rw [Set.infinite_iff_exists_gt] + intro n + obtain ⟨p, ⟨_, h2⟩⟩ := exists_infinite_primes (max (n + 1) 3) + exact ⟨p, Set.mem_setOf.mpr ⟨Prime.odd_of_ne_two h2 (Ne.symm (ne_of_lt (by omega))), + Prime.deficient h2⟩, by omega⟩ + +end Nat diff --git a/docs/references.bib b/docs/references.bib index 71178e6a2c7f3..65985063312cd 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2827,6 +2827,17 @@ @Misc{ ponton2020chebyshev primaryclass = {math.NT} } +@Article{ Prielipp1970, + author = {Robert W. Prielipp}, + title = {PERFECT NUMBERS, ABUNDANT NUMBERS, AND DEFICIENT NUMBERS}, + journal = {The Mathematics Teacher}, + volume = {63}, + year = {1970}, + pages = {692--696}, + issn = {00255769}, + url = {http://www.jstor.org/stable/27958492} +} + @InCollection{ ribenboim1971, author = {Ribenboim, Paulo}, title = {\'{E}pimorphismes de modules qui sont n\'{e}cessairement From 0668f239012afb2e3e10ba2770c7bf9b18f386b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Sep 2024 00:04:20 +0000 Subject: [PATCH 027/174] chore: Rename `DenseInducing` to `IsDenseInducing` (#17208) `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 `DenseInducing` 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/Analysis/InnerProductSpace/Basic.lean | 8 +- .../Analysis/Normed/Group/HomCompletion.lean | 4 +- .../OperatorNorm/Completeness.lean | 6 +- .../Function/SimpleFuncDenseLp.lean | 8 +- Mathlib/Topology/Algebra/GroupCompletion.lean | 4 +- .../Algebra/InfiniteSum/GroupCompletion.lean | 4 +- .../Algebra/Nonarchimedean/Completion.lean | 2 +- Mathlib/Topology/Algebra/UniformField.lean | 12 +-- Mathlib/Topology/Algebra/UniformGroup.lean | 12 +-- Mathlib/Topology/Algebra/UniformRing.lean | 20 ++--- .../Topology/Algebra/Valued/ValuedField.lean | 14 ++-- .../Topology/Compactification/OnePoint.lean | 2 +- Mathlib/Topology/DenseEmbedding.lean | 76 +++++++++---------- Mathlib/Topology/ExtendFrom.lean | 2 +- Mathlib/Topology/Instances/Complex.lean | 10 +-- Mathlib/Topology/Instances/RatLemmas.lean | 2 +- Mathlib/Topology/StoneCech.lean | 14 ++-- .../UniformSpace/AbstractCompletion.lean | 22 +++--- .../UniformSpace/CompleteSeparated.lean | 6 +- Mathlib/Topology/UniformSpace/Completion.lean | 14 ++-- .../UniformSpace/UniformEmbedding.lean | 18 ++--- docs/overview.yaml | 2 +- scripts/no_lints_prime_decls.txt | 4 +- 23 files changed, 133 insertions(+), 133 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 07b4bdc761b64..29e41c81cd70d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -2361,11 +2361,11 @@ open UniformSpace Function instance toInner {𝕜' E' : Type*} [TopologicalSpace 𝕜'] [UniformSpace E'] [Inner 𝕜' E'] : Inner 𝕜' (Completion E') where - inner := curry <| (denseInducing_coe.prod denseInducing_coe).extend (uncurry inner) + inner := curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend (uncurry inner) @[simp] theorem inner_coe (a b : E) : inner (a : Completion E) (b : Completion E) = (inner a b : 𝕜) := - (denseInducing_coe.prod denseInducing_coe).extend_eq + (isDenseInducing_coe.prod isDenseInducing_coe).extend_eq (continuous_inner : Continuous (uncurry inner : E × E → 𝕜)) (a, b) protected theorem continuous_inner : @@ -2378,9 +2378,9 @@ protected theorem continuous_inner : rw [Completion.toInner, inner, uncurry_curry _] change Continuous - (((denseInducing_toCompl E).prod (denseInducing_toCompl E)).extend fun p : E × E => + (((isDenseInducing_toCompl E).prod (isDenseInducing_toCompl E)).extend fun p : E × E => inner' p.1 p.2) - exact (denseInducing_toCompl E).extend_Z_bilin (denseInducing_toCompl E) this + exact (isDenseInducing_toCompl E).extend_Z_bilin (isDenseInducing_toCompl E) this protected theorem Continuous.inner {α : Type*} [TopologicalSpace α] {f g : α → Completion E} (hf : Continuous f) (hg : Continuous g) : Continuous (fun x : α => inner (f x) (g x) : α → 𝕜) := diff --git a/Mathlib/Analysis/Normed/Group/HomCompletion.lean b/Mathlib/Analysis/Normed/Group/HomCompletion.lean index ca780e00bbd2c..3ef2814acd733 100644 --- a/Mathlib/Analysis/Normed/Group/HomCompletion.lean +++ b/Mathlib/Analysis/Normed/Group/HomCompletion.lean @@ -132,7 +132,7 @@ theorem NormedAddCommGroup.norm_toCompl (x : G) : ‖toCompl x‖ = ‖x‖ := Completion.norm_coe x theorem NormedAddCommGroup.denseRange_toCompl : DenseRange (toCompl : G → Completion G) := - Completion.denseInducing_coe.dense + Completion.isDenseInducing_coe.dense @[simp] theorem NormedAddGroupHom.completion_toCompl (f : NormedAddGroupHom G H) : @@ -158,7 +158,7 @@ theorem NormedAddGroupHom.ker_completion {f : NormedAddGroupHom G H} {C : ℝ} rcases h.exists_pos with ⟨C', C'_pos, hC'⟩ rcases exists_pos_mul_lt ε_pos (1 + C' * ‖f‖) with ⟨δ, δ_pos, hδ⟩ obtain ⟨_, ⟨g : G, rfl⟩, hg : ‖hatg - g‖ < δ⟩ := - SeminormedAddCommGroup.mem_closure_iff.mp (Completion.denseInducing_coe.dense hatg) δ δ_pos + SeminormedAddCommGroup.mem_closure_iff.mp (Completion.isDenseInducing_coe.dense hatg) δ δ_pos obtain ⟨g' : G, hgg' : f g' = f g, hfg : ‖g'‖ ≤ C' * ‖f g‖⟩ := hC' (f g) (mem_range_self _ g) have mem_ker : g - g' ∈ f.ker := by rw [f.mem_ker, map_sub, sub_eq_zero.mpr hgg'.symm] refine ⟨_, ⟨⟨g - g', mem_ker⟩, rfl⟩, ?_⟩ diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index b4e87ed65eb32..716b0f8e64316 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -190,7 +190,7 @@ def extend : Fₗ →SL[σ₁₂] F := have cont := (uniformContinuous_uniformly_extend h_e h_dense f.uniformContinuous).continuous -- extension of `f` agrees with `f` on the domain of the embedding `e` have eq := uniformly_extend_of_ind h_e h_dense f.uniformContinuous - { toFun := (h_e.denseInducing h_dense).extend f + { toFun := (h_e.isDenseInducing h_dense).extend f map_add' := by refine h_dense.induction_on₂ ?_ ?_ · exact isClosed_eq (cont.comp continuous_add) @@ -208,10 +208,10 @@ def extend : Fₗ →SL[σ₁₂] F := exact ContinuousLinearMap.map_smulₛₗ _ _ _ cont } --- Porting note: previously `(h_e.denseInducing h_dense)` was inferred. +-- Porting note: previously `(h_e.isDenseInducing h_dense)` was inferred. @[simp] theorem extend_eq (x : E) : extend f e h_dense h_e (e x) = f x := - DenseInducing.extend_eq (h_e.denseInducing h_dense) f.cont _ + IsDenseInducing.extend_eq (h_e.isDenseInducing h_dense) f.cont _ theorem extend_unique (g : Fₗ →SL[σ₁₂] F) (H : g.comp e = f) : extend f e h_dense h_e = g := ContinuousLinearMap.coeFn_injective <| diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 0386ca8b2e120..d2aa0a504c720 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -703,13 +703,13 @@ protected theorem denseEmbedding (hp_ne_top : p ≠ ∞) : convert SimpleFunc.tendsto_approxOn_range_Lp hp_ne_top (Lp.stronglyMeasurable f).measurable hfi' rw [toLp_coeFn f (Lp.memℒp f)] -protected theorem denseInducing (hp_ne_top : p ≠ ∞) : - DenseInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - (simpleFunc.denseEmbedding hp_ne_top).toDenseInducing +protected theorem isDenseInducing (hp_ne_top : p ≠ ∞) : + IsDenseInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := + (simpleFunc.denseEmbedding hp_ne_top).toIsDenseInducing protected theorem denseRange (hp_ne_top : p ≠ ∞) : DenseRange ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - (simpleFunc.denseInducing hp_ne_top).dense + (simpleFunc.isDenseInducing hp_ne_top).dense protected theorem dense (hp_ne_top : p ≠ ∞) : Dense (Lp.simpleFunc E p μ : Set (Lp E p μ)) := by simpa only [denseRange_subtype_val] using simpleFunc.denseRange (E := E) (μ := μ) hp_ne_top diff --git a/Mathlib/Topology/Algebra/GroupCompletion.lean b/Mathlib/Topology/Algebra/GroupCompletion.lean index 234486fbb1753..cadc887e6ae6e 100644 --- a/Mathlib/Topology/Algebra/GroupCompletion.lean +++ b/Mathlib/Topology/Algebra/GroupCompletion.lean @@ -181,8 +181,8 @@ theorem continuous_toCompl : Continuous (toCompl : α → Completion α) := variable (α) -theorem denseInducing_toCompl : DenseInducing (toCompl : α → Completion α) := - denseInducing_coe +theorem isDenseInducing_toCompl : IsDenseInducing (toCompl : α → Completion α) := + isDenseInducing_coe variable {α} diff --git a/Mathlib/Topology/Algebra/InfiniteSum/GroupCompletion.lean b/Mathlib/Topology/Algebra/InfiniteSum/GroupCompletion.lean index 52b2b087f1a48..078edbe10ffe4 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/GroupCompletion.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/GroupCompletion.lean @@ -17,13 +17,13 @@ variable {α β : Type*} [AddCommGroup α] [UniformSpace α] [UniformAddGroup α /-- A function `f` has a sum in an uniform additive group `α` if and only if it has that sum in the completion of `α`. -/ theorem hasSum_iff_hasSum_compl (f : β → α) (a : α) : - HasSum (toCompl ∘ f) a ↔ HasSum f a := (denseInducing_toCompl α).hasSum_iff f a + HasSum (toCompl ∘ f) a ↔ HasSum f a := (isDenseInducing_toCompl α).hasSum_iff f a /-- A function `f` is summable in a uniform additive group `α` if and only if it is summable in `Completion α` and its sum in `Completion α` lies in the range of `toCompl : α →+ Completion α`. -/ theorem summable_iff_summable_compl_and_tsum_mem (f : β → α) : Summable f ↔ Summable (toCompl ∘ f) ∧ ∑' i, toCompl (f i) ∈ Set.range toCompl := - (denseInducing_toCompl α).summable_iff_tsum_comp_mem_range f + (isDenseInducing_toCompl α).summable_iff_tsum_comp_mem_range f /-- A function `f` is summable in a uniform additive group `α` if and only if the net of its partial sums is Cauchy and its sum in `Completion α` lies in the range of `toCompl : α →+ Completion α`. diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Completion.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Completion.lean index 8d88cf4ee0a0a..c5dd61852fecd 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Completion.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Completion.lean @@ -51,7 +51,7 @@ instance {G : Type*} [AddGroup G] [UniformSpace G] [UniformAddGroup G] [Nonarchi `0` in `Completion G`. This follows from the fact that `toCompl : G → Completion G` is dense inducing and `W` is a neighborhood of `0` in `G`. -/ apply isOpen_of_mem_nhds (g := 0) - apply (denseInducing_toCompl _).closure_image_mem_nhds + apply (isDenseInducing_toCompl _).closure_image_mem_nhds exact mem_nhds_zero W use ⟨_, this⟩ /- Finally, it remains to show that `V ⊆ U`. It suffices to show that `V ⊆ C`, which diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 059b0ae0cacfb..59733a4968e23 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -61,11 +61,11 @@ variable {K} /-- extension of inversion to the completion of a field. -/ def hatInv : hat K → hat K := - denseInducing_coe.extend fun x : K => (↑x⁻¹ : hat K) + isDenseInducing_coe.extend fun x : K => (↑x⁻¹ : hat K) theorem continuous_hatInv [CompletableTopField K] {x : hat K} (h : x ≠ 0) : ContinuousAt hatInv x := by - refine denseInducing_coe.continuousAt_extend ?_ + refine isDenseInducing_coe.continuousAt_extend ?_ apply mem_of_superset (compl_singleton_mem_nhds h) intro y y_ne rw [mem_compl_singleton_iff] at y_ne @@ -77,13 +77,13 @@ theorem continuous_hatInv [CompletableTopField K] {x : hat K} (h : x ≠ 0) : rw [this, ← Filter.map_map] apply Cauchy.map _ (Completion.uniformContinuous_coe K) apply CompletableTopField.nice - · haveI := denseInducing_coe.comap_nhds_neBot y + · haveI := isDenseInducing_coe.comap_nhds_neBot y apply cauchy_nhds.comap rw [Completion.comap_coe_eq_uniformity] · have eq_bot : 𝓝 (0 : hat K) ⊓ 𝓝 y = ⊥ := by by_contra h exact y_ne (eq_of_nhds_neBot <| neBot_iff.mpr h).symm - erw [denseInducing_coe.nhds_eq_comap (0 : K), ← Filter.comap_inf, eq_bot] + erw [isDenseInducing_coe.nhds_eq_comap (0 : K), ← Filter.comap_inf, eq_bot] exact comap_bot open Classical in @@ -97,7 +97,7 @@ instance instInvCompletion : Inv (hat K) := variable [TopologicalDivisionRing K] theorem hatInv_extends {x : K} (h : x ≠ 0) : hatInv (x : hat K) = ↑(x⁻¹ : K) := - denseInducing_coe.extend_eq_at ((continuous_coe K).continuousAt.comp (continuousAt_inv₀ h)) + isDenseInducing_coe.extend_eq_at ((continuous_coe K).continuousAt.comp (continuousAt_inv₀ h)) variable [CompletableTopField K] @@ -126,7 +126,7 @@ theorem mul_hatInv_cancel {x : hat K} (x_ne : x ≠ 0) : x * hatInv x = 1 := by continuous_id.continuousAt.prod (continuous_hatInv x_ne) exact (_root_.continuous_mul.continuousAt.comp this : _) have clo : x ∈ closure (c '' {0}ᶜ) := by - have := denseInducing_coe.dense x + have := isDenseInducing_coe.dense x rw [← image_univ, show (univ : Set K) = {0} ∪ {0}ᶜ from (union_compl_self _).symm, image_union] at this apply mem_closure_of_mem_closure_union this diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index f6f8cdc90796d..21f75fd7a2065 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -624,7 +624,7 @@ variable [TopologicalSpace β] [Group β] variable [FunLike hom β α] [MonoidHomClass hom β α] {e : hom} @[to_additive] -theorem tendsto_div_comap_self (de : DenseInducing e) (x₀ : α) : +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)) = @@ -638,7 +638,7 @@ theorem tendsto_div_comap_self (de : DenseInducing e) (x₀ : α) : end -namespace DenseInducing +namespace IsDenseInducing variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} variable {G : Type*} @@ -650,8 +650,8 @@ variable [TopologicalSpace β] [AddCommGroup β] variable [TopologicalSpace γ] [AddCommGroup γ] [TopologicalAddGroup γ] variable [TopologicalSpace δ] [AddCommGroup δ] variable [UniformSpace G] [AddCommGroup G] -variable {e : β →+ α} (de : DenseInducing e) -variable {f : δ →+ γ} (df : DenseInducing f) +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)) @@ -725,7 +725,7 @@ private theorem extend_Z_bilin_key (x₀ : α) (y₀ : γ) : ∃ U ∈ comap e ( have h₄ := H x₁ x₁_in x xU₁ y yV₁ y' y'V₁ exact W4 h₁ h₂ h₃ h₄ -open DenseInducing +open IsDenseInducing variable [T0Space G] [CompleteSpace G] @@ -767,7 +767,7 @@ theorem extend_Z_bilin : Continuous (extend (de.prod df) (fun p : β × δ => φ rcases p with ⟨⟨x, y⟩, ⟨x', y'⟩⟩ apply h <;> tauto -end DenseInducing +end IsDenseInducing section CompleteQuotient diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index d5096b53d82a7..f4799994dceaf 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -38,7 +38,7 @@ noncomputable section universe u namespace UniformSpace.Completion -open DenseInducing UniformSpace Function +open IsDenseInducing UniformSpace Function section one_and_mul variable (α : Type*) [Ring α] [UniformSpace α] @@ -47,7 +47,7 @@ instance one : One (Completion α) := ⟨(1 : α)⟩ instance mul : Mul (Completion α) := - ⟨curry <| (denseInducing_coe.prod denseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ + ⟨curry <| (isDenseInducing_coe.prod isDenseInducing_coe).extend ((↑) ∘ uncurry (· * ·))⟩ @[norm_cast] theorem coe_one : ((1 : α) : Completion α) = 1 := @@ -59,7 +59,7 @@ variable {α : Type*} [Ring α] [UniformSpace α] [TopologicalRing α] @[norm_cast] theorem coe_mul (a b : α) : ((a * b : α) : Completion α) = a * b := - ((denseInducing_coe.prod denseInducing_coe).extend_eq + ((isDenseInducing_coe.prod isDenseInducing_coe).extend_eq ((continuous_coe α).comp (@continuous_mul α _ _ _)) (a, b)).symm variable [UniformAddGroup α] @@ -70,7 +70,7 @@ theorem continuous_mul : Continuous fun p : Completion α × Completion α => p. apply (continuous_coe α).comp _ simp only [AddMonoidHom.coe_mul, AddMonoidHom.coe_mulLeft] exact _root_.continuous_mul - have di : DenseInducing (toCompl : α → Completion α) := denseInducing_coe + have di : IsDenseInducing (toCompl : α → Completion α) := isDenseInducing_coe convert di.extend_Z_bilin di this theorem Continuous.mul {β : Type*} [TopologicalSpace β] {f g : β → Completion α} @@ -276,28 +276,28 @@ variable {γ : Type*} [UniformSpace γ] [Semiring γ] [TopologicalSemiring γ] variable [T2Space γ] [CompleteSpace γ] /-- The dense inducing extension as a ring homomorphism. -/ -noncomputable def DenseInducing.extendRingHom {i : α →+* β} {f : α →+* γ} (ue : UniformInducing i) +noncomputable def IsDenseInducing.extendRingHom {i : α →+* β} {f : α →+* γ} (ue : UniformInducing i) (dr : DenseRange i) (hf : UniformContinuous f) : β →+* γ where - toFun := (ue.denseInducing dr).extend f + toFun := (ue.isDenseInducing dr).extend f map_one' := by - convert DenseInducing.extend_eq (ue.denseInducing dr) hf.continuous 1 + convert IsDenseInducing.extend_eq (ue.isDenseInducing dr) hf.continuous 1 exacts [i.map_one.symm, f.map_one.symm] map_zero' := by - convert DenseInducing.extend_eq (ue.denseInducing dr) hf.continuous 0 <;> + convert IsDenseInducing.extend_eq (ue.isDenseInducing dr) hf.continuous 0 <;> simp only [map_zero] map_add' := by have h := (uniformContinuous_uniformly_extend ue dr hf).continuous refine fun x y => DenseRange.induction_on₂ dr ?_ (fun a b => ?_) x y · exact isClosed_eq (Continuous.comp h continuous_add) ((h.comp continuous_fst).add (h.comp continuous_snd)) - · simp_rw [← i.map_add, DenseInducing.extend_eq (ue.denseInducing dr) hf.continuous _, + · simp_rw [← i.map_add, IsDenseInducing.extend_eq (ue.isDenseInducing dr) hf.continuous _, ← f.map_add] map_mul' := by have h := (uniformContinuous_uniformly_extend ue dr hf).continuous refine fun x y => DenseRange.induction_on₂ dr ?_ (fun a b => ?_) x y · exact isClosed_eq (Continuous.comp h continuous_mul) ((h.comp continuous_fst).mul (h.comp continuous_snd)) - · simp_rw [← i.map_mul, DenseInducing.extend_eq (ue.denseInducing dr) hf.continuous _, + · simp_rw [← i.map_mul, IsDenseInducing.extend_eq (ue.isDenseInducing dr) hf.continuous _, ← f.map_mul] end UniformExtension diff --git a/Mathlib/Topology/Algebra/Valued/ValuedField.lean b/Mathlib/Topology/Algebra/Valued/ValuedField.lean index c9a159cdcd849..99151442537d6 100644 --- a/Mathlib/Topology/Algebra/Valued/ValuedField.lean +++ b/Mathlib/Topology/Algebra/Valued/ValuedField.lean @@ -187,14 +187,14 @@ open WithZeroTopology /-- The extension of the valuation of a valued field to the completion of the field. -/ noncomputable def extension : hat K → Γ₀ := - Completion.denseInducing_coe.extend (v : K → Γ₀) + Completion.isDenseInducing_coe.extend (v : K → Γ₀) theorem continuous_extension : Continuous (Valued.extension : hat K → Γ₀) := by - refine Completion.denseInducing_coe.continuous_extend ?_ + refine Completion.isDenseInducing_coe.continuous_extend ?_ intro x₀ rcases eq_or_ne x₀ 0 with (rfl | h) · refine ⟨0, ?_⟩ - erw [← Completion.denseInducing_coe.toInducing.nhds_eq_comap] + erw [← Completion.isDenseInducing_coe.toInducing.nhds_eq_comap] exact Valued.continuous_valuation.tendsto' 0 0 (map_zero v) · have preimage_one : v ⁻¹' {(1 : Γ₀)} ∈ 𝓝 (1 : K) := by have : (v (1 : K) : Γ₀) ≠ 0 := by @@ -204,7 +204,7 @@ theorem continuous_extension : Continuous (Valued.extension : hat K → Γ₀) : ext x rw [Valuation.map_one, mem_preimage, mem_singleton_iff, mem_setOf_eq] obtain ⟨V, V_in, hV⟩ : ∃ V ∈ 𝓝 (1 : hat K), ∀ x : K, (x : hat K) ∈ V → (v x : Γ₀) = 1 := by - rwa [Completion.denseInducing_coe.nhds_eq_comap, mem_comap] at preimage_one + rwa [Completion.isDenseInducing_coe.nhds_eq_comap, mem_comap] at preimage_one have : ∃ V' ∈ 𝓝 (1 : hat K), (0 : hat K) ∉ V' ∧ ∀ (x) (_ : x ∈ V') (y) (_ : y ∈ V'), x * y⁻¹ ∈ V := by have : Tendsto (fun p : hat K × hat K => p.1 * p.2⁻¹) ((𝓝 1) ×ˢ (𝓝 1)) (𝓝 1) := by @@ -265,8 +265,8 @@ theorem continuous_extension : Continuous (Valued.extension : hat K → Γ₀) : @[simp, norm_cast] theorem extension_extends (x : K) : extension (x : hat K) = v x := by - refine Completion.denseInducing_coe.extend_eq_of_tendsto ?_ - rw [← Completion.denseInducing_coe.nhds_eq_comap] + refine Completion.isDenseInducing_coe.extend_eq_of_tendsto ?_ + rw [← Completion.isDenseInducing_coe.nhds_eq_comap] exact Valued.continuous_valuation.continuousAt /-- the extension of a valuation on a division ring to its completion. -/ @@ -339,7 +339,7 @@ noncomputable instance valuedCompletion : Valued (hat K) Γ₀ where rw [this.mem_iff] exact exists_congr fun γ => by simp simp_rw [← closure_coe_completion_v_lt] - exact (hasBasis_nhds_zero K Γ₀).hasBasis_of_denseInducing Completion.denseInducing_coe + exact (hasBasis_nhds_zero K Γ₀).hasBasis_of_isDenseInducing Completion.isDenseInducing_coe -- Porting note: removed @[norm_cast] attribute due to error: -- norm_cast: badly shaped lemma, rhs can't start with coe diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index 0201f2ccbee78..cb7d3c40f534c 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -507,7 +507,7 @@ example [WeaklyLocallyCompactSpace X] [T2Space X] : T4Space (OnePoint X) := infe /-- If `X` is not a compact space, then `OnePoint X` is a connected space. -/ instance [PreconnectedSpace X] [NoncompactSpace X] : ConnectedSpace (OnePoint X) where - toPreconnectedSpace := denseEmbedding_coe.toDenseInducing.preconnectedSpace + toPreconnectedSpace := denseEmbedding_coe.toIsDenseInducing.preconnectedSpace toNonempty := inferInstance /-- If `X` is an infinite type with discrete topology (e.g., `ℕ`), then the identity map from diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index f346261d50bed..78c6173ce6866 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -12,13 +12,13 @@ import Mathlib.Topology.Bases This file defines three properties of functions: * `DenseRange f` means `f` has dense image; -* `DenseInducing i` means `i` is also `Inducing`, namely it induces the topology on its codomain; +* `IsDenseInducing i` means `i` is also `Inducing`, namely it induces the topology on its codomain; * `DenseEmbedding e` means `e` is further an `Embedding`, namely it is injective and `Inducing`. The main theorem `continuous_extend` gives a criterion for a function `f : X → Z` to a T₃ space Z to extend along a dense embedding `i : X → Y` to a continuous function `g : Y → Z`. Actually `i` only -has to be `DenseInducing` (not necessarily injective). +has to be `IsDenseInducing` (not necessarily injective). -/ @@ -32,30 +32,30 @@ variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} /-- `i : α → β` is "dense inducing" if it has dense range and the topology on `α` is the one induced by `i` from the topology on `β`. -/ -structure DenseInducing [TopologicalSpace α] [TopologicalSpace β] (i : α → β) +structure IsDenseInducing [TopologicalSpace α] [TopologicalSpace β] (i : α → β) extends Inducing i : Prop where /-- The range of a dense inducing map is a dense set. -/ protected dense : DenseRange i -namespace DenseInducing +namespace IsDenseInducing variable [TopologicalSpace α] [TopologicalSpace β] variable {i : α → β} -theorem nhds_eq_comap (di : DenseInducing i) : ∀ a : α, 𝓝 a = comap i (𝓝 <| i a) := +theorem nhds_eq_comap (di : IsDenseInducing i) : ∀ a : α, 𝓝 a = comap i (𝓝 <| i a) := di.toInducing.nhds_eq_comap -protected theorem continuous (di : DenseInducing i) : Continuous i := +protected theorem continuous (di : IsDenseInducing i) : Continuous i := di.toInducing.continuous -theorem closure_range (di : DenseInducing i) : closure (range i) = univ := +theorem closure_range (di : IsDenseInducing i) : closure (range i) = univ := di.dense.closure_range -protected theorem preconnectedSpace [PreconnectedSpace α] (di : DenseInducing i) : +protected theorem preconnectedSpace [PreconnectedSpace α] (di : IsDenseInducing i) : PreconnectedSpace β := di.dense.preconnectedSpace di.continuous -theorem closure_image_mem_nhds {s : Set α} {a : α} (di : DenseInducing i) (hs : s ∈ 𝓝 a) : +theorem closure_image_mem_nhds {s : Set α} {a : α} (di : IsDenseInducing i) (hs : s ∈ 𝓝 a) : closure (i '' s) ∈ 𝓝 (i a) := by rw [di.nhds_eq_comap a, ((nhds_basis_opens _).comap _).mem_iff] at hs rcases hs with ⟨U, ⟨haU, hUo⟩, sub : i ⁻¹' U ⊆ s⟩ @@ -64,14 +64,14 @@ theorem closure_image_mem_nhds {s : Set α} {a : α} (di : DenseInducing i) (hs U ⊆ closure (i '' (i ⁻¹' U)) := di.dense.subset_closure_image_preimage_of_isOpen hUo _ ⊆ closure (i '' s) := closure_mono (image_subset i sub) -theorem dense_image (di : DenseInducing i) {s : Set α} : Dense (i '' s) ↔ Dense s := by +theorem dense_image (di : IsDenseInducing i) {s : Set α} : Dense (i '' s) ↔ Dense s := by refine ⟨fun H x => ?_, di.dense.dense_image di.continuous⟩ rw [di.toInducing.closure_eq_preimage_closure_image, H.closure_eq, preimage_univ] trivial /-- If `i : α → β` is a dense embedding with dense complement of the range, then any compact set in `α` has empty interior. -/ -theorem interior_compact_eq_empty [T2Space β] (di : DenseInducing i) (hd : Dense (range i)ᶜ) +theorem interior_compact_eq_empty [T2Space β] (di : IsDenseInducing i) (hd : Dense (range i)ᶜ) {s : Set α} (hs : IsCompact s) : interior s = ∅ := by refine eq_empty_iff_forall_not_mem.2 fun x hx => ?_ rw [mem_interior_iff_mem_nhds] at hx @@ -82,15 +82,15 @@ theorem interior_compact_eq_empty [T2Space β] (di : DenseInducing i) (hd : Dens /-- The product of two dense inducings is a dense inducing -/ protected theorem prod [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α → β} {e₂ : γ → δ} - (de₁ : DenseInducing e₁) (de₂ : DenseInducing e₂) : - DenseInducing fun p : α × γ => (e₁ p.1, e₂ p.2) where + (de₁ : IsDenseInducing e₁) (de₂ : IsDenseInducing e₂) : + IsDenseInducing fun p : α × γ => (e₁ p.1, e₂ p.2) where toInducing := de₁.toInducing.prod_map de₂.toInducing dense := de₁.dense.prod_map de₂.dense open TopologicalSpace -/-- If the domain of a `DenseInducing` map is a separable space, then so is the codomain. -/ -protected theorem separableSpace [SeparableSpace α] (di : DenseInducing i) : SeparableSpace β := +/-- If the domain of a `IsDenseInducing` map is a separable space, then so is the codomain. -/ +protected theorem separableSpace [SeparableSpace α] (di : IsDenseInducing i) : SeparableSpace β := di.dense.separableSpace di.continuous variable [TopologicalSpace δ] {f : γ → α} {g : γ → δ} {h : δ → β} @@ -102,7 +102,7 @@ g↓ ↓e δ -h→ β ``` -/ -theorem tendsto_comap_nhds_nhds {d : δ} {a : α} (di : DenseInducing i) +theorem tendsto_comap_nhds_nhds {d : δ} {a : α} (di : IsDenseInducing i) (H : Tendsto h (𝓝 d) (𝓝 (i a))) (comm : h ∘ g = i ∘ f) : Tendsto f (comap g (𝓝 d)) (𝓝 a) := by have lim1 : map g (comap g (𝓝 d)) ≤ 𝓝 d := map_comap_le replace lim1 : map h (map g (comap g (𝓝 d))) ≤ map h (𝓝 d) := map_mono lim1 @@ -111,10 +111,10 @@ theorem tendsto_comap_nhds_nhds {d : δ} {a : α} (di : DenseInducing i) rw [← di.nhds_eq_comap] at lim2 exact le_trans lim1 lim2 -protected theorem nhdsWithin_neBot (di : DenseInducing i) (b : β) : NeBot (𝓝[range i] b) := +protected theorem nhdsWithin_neBot (di : IsDenseInducing i) (b : β) : NeBot (𝓝[range i] b) := di.dense.nhdsWithin_neBot b -theorem comap_nhds_neBot (di : DenseInducing i) (b : β) : NeBot (comap i (𝓝 b)) := +theorem comap_nhds_neBot (di : IsDenseInducing i) (b : β) : NeBot (comap i (𝓝 b)) := comap_neBot fun s hs => by rcases mem_closure_iff_nhds.1 (di.dense b) s hs with ⟨_, ⟨ha, a, rfl⟩⟩ exact ⟨a, ha⟩ @@ -122,38 +122,38 @@ theorem comap_nhds_neBot (di : DenseInducing i) (b : β) : NeBot (comap i (𝓝 variable [TopologicalSpace γ] /-- If `i : α → β` is a dense inducing, then any function `f : α → γ` "extends" to a function `g = - DenseInducing.extend di f : β → γ`. If `γ` is Hausdorff and `f` has a continuous extension, then + IsDenseInducing.extend di f : β → γ`. If `γ` is Hausdorff and `f` has a continuous extension, then `g` is the unique such extension. In general, `g` might not be continuous or even extend `f`. -/ -def extend (di : DenseInducing i) (f : α → γ) (b : β) : γ := +def extend (di : IsDenseInducing i) (f : α → γ) (b : β) : γ := @limUnder _ _ _ ⟨f (di.dense.some b)⟩ (comap i (𝓝 b)) f -theorem extend_eq_of_tendsto [T2Space γ] (di : DenseInducing i) {b : β} {c : γ} {f : α → γ} +theorem extend_eq_of_tendsto [T2Space γ] (di : IsDenseInducing i) {b : β} {c : γ} {f : α → γ} (hf : Tendsto f (comap i (𝓝 b)) (𝓝 c)) : di.extend f b = c := haveI := di.comap_nhds_neBot hf.limUnder_eq -theorem extend_eq_at [T2Space γ] (di : DenseInducing i) {f : α → γ} {a : α} +theorem extend_eq_at [T2Space γ] (di : IsDenseInducing i) {f : α → γ} {a : α} (hf : ContinuousAt f a) : di.extend f (i a) = f a := extend_eq_of_tendsto _ <| di.nhds_eq_comap a ▸ hf -theorem extend_eq_at' [T2Space γ] (di : DenseInducing i) {f : α → γ} {a : α} (c : γ) +theorem extend_eq_at' [T2Space γ] (di : IsDenseInducing i) {f : α → γ} {a : α} (c : γ) (hf : Tendsto f (𝓝 a) (𝓝 c)) : di.extend f (i a) = f a := di.extend_eq_at (continuousAt_of_tendsto_nhds hf) -theorem extend_eq [T2Space γ] (di : DenseInducing i) {f : α → γ} (hf : Continuous f) (a : α) : +theorem extend_eq [T2Space γ] (di : IsDenseInducing i) {f : α → γ} (hf : Continuous f) (a : α) : di.extend f (i a) = f a := di.extend_eq_at hf.continuousAt /-- Variation of `extend_eq` where we ask that `f` has a limit along `comap i (𝓝 b)` for each `b : β`. This is a strictly stronger assumption than continuity of `f`, but in a lot of cases you'd have to prove it anyway to use `continuous_extend`, so this avoids doing the work twice. -/ -theorem extend_eq' [T2Space γ] {f : α → γ} (di : DenseInducing i) +theorem extend_eq' [T2Space γ] {f : α → γ} (di : IsDenseInducing i) (hf : ∀ b, ∃ c, Tendsto f (comap i (𝓝 b)) (𝓝 c)) (a : α) : di.extend f (i a) = f a := by rcases hf (i a) with ⟨b, hb⟩ refine di.extend_eq_at' b ?_ rwa [← di.toInducing.nhds_eq_comap] at hb -theorem extend_unique_at [T2Space γ] {b : β} {f : α → γ} {g : β → γ} (di : DenseInducing i) +theorem extend_unique_at [T2Space γ] {b : β} {f : α → γ} {g : β → γ} (di : IsDenseInducing i) (hf : ∀ᶠ x in comap i (𝓝 b), g (i x) = f x) (hg : ContinuousAt g b) : di.extend f b = g b := by refine di.extend_eq_of_tendsto fun s hs => mem_map.2 ?_ suffices ∀ᶠ x : α in comap i (𝓝 b), g (i x) ∈ s from @@ -163,11 +163,11 @@ theorem extend_unique_at [T2Space γ] {b : β} {f : α → γ} {g : β → γ} ( rintro _ hxs x rfl exact hxs -theorem extend_unique [T2Space γ] {f : α → γ} {g : β → γ} (di : DenseInducing i) +theorem extend_unique [T2Space γ] {f : α → γ} {g : β → γ} (di : IsDenseInducing i) (hf : ∀ x, g (i x) = f x) (hg : Continuous g) : di.extend f = g := funext fun _ => extend_unique_at di (Eventually.of_forall hf) hg.continuousAt -theorem continuousAt_extend [T3Space γ] {b : β} {f : α → γ} (di : DenseInducing i) +theorem continuousAt_extend [T3Space γ] {b : β} {f : α → γ} (di : IsDenseInducing i) (hf : ∀ᶠ x in 𝓝 b, ∃ c, Tendsto f (comap i <| 𝓝 x) (𝓝 c)) : ContinuousAt (di.extend f) b := by set φ := di.extend f haveI := di.comap_nhds_neBot @@ -189,28 +189,28 @@ theorem continuousAt_extend [T3Space γ] {b : β} {f : α → γ} (di : DenseInd use V₂ tauto -theorem continuous_extend [T3Space γ] {f : α → γ} (di : DenseInducing i) +theorem continuous_extend [T3Space γ] {f : α → γ} (di : IsDenseInducing i) (hf : ∀ b, ∃ c, Tendsto f (comap i (𝓝 b)) (𝓝 c)) : Continuous (di.extend f) := continuous_iff_continuousAt.mpr fun _ => di.continuousAt_extend <| univ_mem' hf theorem mk' (i : α → β) (c : Continuous i) (dense : ∀ x, x ∈ closure (range i)) - (H : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 (i a), ∀ b, i b ∈ t → b ∈ s) : DenseInducing i := + (H : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 (i a), ∀ b, i b ∈ t → b ∈ s) : IsDenseInducing i := { toInducing := inducing_iff_nhds.2 fun a => le_antisymm (c.tendsto _).le_comap (by simpa [Filter.le_def] using H a) dense } -end DenseInducing +end IsDenseInducing /-- A dense embedding is an embedding with dense image. -/ structure DenseEmbedding [TopologicalSpace α] [TopologicalSpace β] (e : α → β) extends - DenseInducing e : Prop where + IsDenseInducing e : Prop where /-- A dense embedding is injective. -/ inj : Function.Injective e theorem DenseEmbedding.mk' [TopologicalSpace α] [TopologicalSpace β] (e : α → β) (c : Continuous e) (dense : DenseRange e) (inj : Function.Injective e) (H : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 (e a), ∀ b, e b ∈ t → b ∈ s) : DenseEmbedding e := - { DenseInducing.mk' e c dense H with inj } + { IsDenseInducing.mk' e c dense H with inj } namespace DenseEmbedding @@ -228,12 +228,12 @@ theorem to_embedding (de : DenseEmbedding e) : Embedding e := /-- If the domain of a `DenseEmbedding` is a separable space, then so is its codomain. -/ protected theorem separableSpace [SeparableSpace α] (de : DenseEmbedding e) : SeparableSpace β := - de.toDenseInducing.separableSpace + de.toIsDenseInducing.separableSpace /-- The product of two dense embeddings is a dense embedding. -/ protected theorem prod {e₁ : α → β} {e₂ : γ → δ} (de₁ : DenseEmbedding e₁) (de₂ : DenseEmbedding e₂) : DenseEmbedding fun p : α × γ => (e₁ p.1, e₂ p.2) := - { de₁.toDenseInducing.prod de₂.toDenseInducing with + { de₁.toIsDenseInducing.prod de₂.toIsDenseInducing with inj := de₁.inj.prodMap de₂.inj } /-- The dense embedding of a subtype inside its closure. -/ @@ -256,7 +256,7 @@ protected theorem subtype (de : DenseEmbedding e) (p : α → Prop) : Function.comp_def] } theorem dense_image (de : DenseEmbedding e) {s : Set α} : Dense (e '' s) ↔ Dense s := - de.toDenseInducing.dense_image + de.toIsDenseInducing.dense_image end DenseEmbedding @@ -319,9 +319,9 @@ theorem DenseRange.equalizer (hfd : DenseRange f) {g h : β → γ} (hg : Contin end -- Bourbaki GT III §3 no.4 Proposition 7 (generalised to any dense-inducing map to a T₃ space) -theorem Filter.HasBasis.hasBasis_of_denseInducing [TopologicalSpace α] [TopologicalSpace β] +theorem Filter.HasBasis.hasBasis_of_isDenseInducing [TopologicalSpace α] [TopologicalSpace β] [T3Space β] {ι : Type*} {s : ι → Set α} {p : ι → Prop} {x : α} (h : (𝓝 x).HasBasis p s) - {f : α → β} (hf : DenseInducing f) : (𝓝 (f x)).HasBasis p fun i => closure <| f '' s i := by + {f : α → β} (hf : IsDenseInducing f) : (𝓝 (f x)).HasBasis p fun i => closure <| f '' s i := by rw [Filter.hasBasis_iff] at h ⊢ intro T refine ⟨fun hT => ?_, fun hT => ?_⟩ diff --git a/Mathlib/Topology/ExtendFrom.lean b/Mathlib/Topology/ExtendFrom.lean index e425d7ce55e1a..1c043eb2756ea 100644 --- a/Mathlib/Topology/ExtendFrom.lean +++ b/Mathlib/Topology/ExtendFrom.lean @@ -12,7 +12,7 @@ The main definition of this file is `extendFrom A f` where `f : X → Y` and `A : Set X`. This defines a new function `g : X → Y` which maps any `x₀ : X` to the limit of `f` as `x` tends to `x₀`, if such a limit exists. -This is analogous to the way `DenseInducing.extend` "extends" a function +This is analogous to the way `IsDenseInducing.extend` "extends" a function `f : X → Z` to a function `g : Y → Z` along a dense inducing `i : X → Y`. The main theorem we prove about this definition is `continuousOn_extendFrom` diff --git a/Mathlib/Topology/Instances/Complex.lean b/Mathlib/Topology/Instances/Complex.lean index 1bd41163d0387..8cfe0e613fcb3 100644 --- a/Mathlib/Topology/Instances/Complex.lean +++ b/Mathlib/Topology/Instances/Complex.lean @@ -56,9 +56,9 @@ theorem Complex.uniformContinuous_ringHom_eq_id_or_conj (K : Subfield ℂ) {ψ : ⟨by erw [uniformity_subtype, uniformity_subtype, Filter.comap_comap] congr ⟩ - let di := ui.denseInducing (?_ : DenseRange ι) + let di := ui.isDenseInducing (?_ : DenseRange ι) · -- extψ : closure(K) →+* ℂ is the extension of ψ : K →+* ℂ - let extψ := DenseInducing.extendRingHom ui di.dense hc + let extψ := IsDenseInducing.extendRingHom ui di.dense hc haveI hψ := (uniformContinuous_uniformly_extend ui di.dense hc).continuous cases' Complex.subfield_eq_of_closed (Subfield.isClosed_topologicalClosure K) with h h · left @@ -76,7 +76,7 @@ theorem Complex.uniformContinuous_ringHom_eq_id_or_conj (K : Subfield ℂ) {ψ : -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [RingHom.comp_apply, RingHom.comp_apply, hr, RingEquiv.toRingHom_eq_coe] at this convert this using 1 - · exact (DenseInducing.extend_eq di hc.continuous _).symm + · exact (IsDenseInducing.extend_eq di hc.continuous _).symm · rw [← ofReal.coe_rangeRestrict, hr] rfl obtain ⟨r, hr⟩ := SetLike.coe_mem (j (ι x)) @@ -94,11 +94,11 @@ theorem Complex.uniformContinuous_ringHom_eq_id_or_conj (K : Subfield ℂ) {ψ : · left ext1 z convert RingHom.congr_fun h z using 1 - exact (DenseInducing.extend_eq di hc.continuous z).symm + exact (IsDenseInducing.extend_eq di hc.continuous z).symm · right ext1 z convert RingHom.congr_fun h z using 1 - exact (DenseInducing.extend_eq di hc.continuous z).symm + exact (IsDenseInducing.extend_eq di hc.continuous z).symm · let j : { x // x ∈ closure (id '' { x | (K : Set ℂ) x }) } → (K.topologicalClosure : Set ℂ) := fun x => ⟨x, by diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index b94351980b416..16c498e04a27f 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -39,7 +39,7 @@ namespace Rat variable {p q : ℚ} {s t : Set ℚ} theorem interior_compact_eq_empty (hs : IsCompact s) : interior s = ∅ := - denseEmbedding_coe_real.toDenseInducing.interior_compact_eq_empty dense_irrational hs + denseEmbedding_coe_real.toIsDenseInducing.interior_compact_eq_empty dense_irrational hs theorem dense_compl_compact (hs : IsCompact s) : Dense sᶜ := interior_eq_empty_iff_dense_compl.1 (interior_compact_eq_empty hs) diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index 34a2c255120e0..240f69c279ec2 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -148,7 +148,7 @@ theorem induced_topology_pure : simp /-- `pure : α → Ultrafilter α` defines a dense inducing of `α` in `Ultrafilter α`. -/ -theorem denseInducing_pure : @DenseInducing _ _ ⊥ _ (pure : α → Ultrafilter α) := +theorem isDenseInducing_pure : @IsDenseInducing _ _ ⊥ _ (pure : α → Ultrafilter α) := letI : TopologicalSpace α := ⊥ ⟨⟨induced_topology_pure.symm⟩, denseRange_pure⟩ @@ -156,7 +156,7 @@ theorem denseInducing_pure : @DenseInducing _ _ ⊥ _ (pure : α → Ultrafilter /-- `pure : α → Ultrafilter α` defines a dense embedding of `α` in `Ultrafilter α`. -/ theorem denseEmbedding_pure : @DenseEmbedding _ _ ⊥ _ (pure : α → Ultrafilter α) := letI : TopologicalSpace α := ⊥ - { denseInducing_pure with inj := ultrafilter_pure_injective } + { isDenseInducing_pure with inj := ultrafilter_pure_injective } end Embedding @@ -166,21 +166,21 @@ section Extension unique extension to a continuous function `Ultrafilter α → γ`. We already know it must be unique because `α → Ultrafilter α` is a dense embedding and `γ` is Hausdorff. For existence, we will invoke - `DenseInducing.continuous_extend`. -/ + `IsDenseInducing.continuous_extend`. -/ variable {γ : Type*} [TopologicalSpace γ] /-- The extension of a function `α → γ` to a function `Ultrafilter α → γ`. When `γ` is a compact Hausdorff space it will be continuous. -/ def Ultrafilter.extend (f : α → γ) : Ultrafilter α → γ := letI : TopologicalSpace α := ⊥ - denseInducing_pure.extend f + isDenseInducing_pure.extend f variable [T2Space γ] theorem ultrafilter_extend_extends (f : α → γ) : Ultrafilter.extend f ∘ pure = f := by letI : TopologicalSpace α := ⊥ haveI : DiscreteTopology α := ⟨rfl⟩ - exact funext (denseInducing_pure.extend_eq continuous_of_discreteTopology) + exact funext (isDenseInducing_pure.extend_eq continuous_of_discreteTopology) variable [CompactSpace γ] @@ -191,7 +191,7 @@ theorem continuous_ultrafilter_extend (f : α → γ) : Continuous (Ultrafilter. isCompact_univ.ultrafilter_le_nhds (b.map f) (by rw [le_principal_iff]; exact univ_mem) ⟨c, le_trans (map_mono (ultrafilter_comap_pure_nhds _)) h'⟩ let _ : TopologicalSpace α := ⊥ - exact denseInducing_pure.continuous_extend h + exact isDenseInducing_pure.continuous_extend h /-- The value of `Ultrafilter.extend f` on an ultrafilter `b` is the unique limit of the ultrafilter `b.map f` in `γ`. -/ @@ -210,7 +210,7 @@ theorem ultrafilter_extend_eq_iff {f : α → γ} {b : Ultrafilter α} {c : γ} exact le_rfl, fun h ↦ let _ : TopologicalSpace α := ⊥ - denseInducing_pure.extend_eq_of_tendsto + isDenseInducing_pure.extend_eq_of_tendsto (le_trans (map_mono (ultrafilter_comap_pure_nhds _)) h)⟩ end Extension diff --git a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean index 903552b0c52dd..9b5cc359931bb 100644 --- a/Mathlib/Topology/UniformSpace/AbstractCompletion.lean +++ b/Mathlib/Topology/UniformSpace/AbstractCompletion.lean @@ -88,7 +88,7 @@ def ofComplete [T0Space α] [CompleteSpace α] : AbstractCompletion α := theorem closure_range : closure (range ι) = univ := pkg.dense.closure_range -theorem denseInducing : DenseInducing ι := +theorem isDenseInducing : IsDenseInducing ι := ⟨pkg.uniformInducing.inducing, pkg.dense⟩ theorem uniformContinuous_coe : UniformContinuous ι := @@ -114,16 +114,16 @@ section Extend /-- Extension of maps to completions -/ protected def extend (f : α → β) : hatα → β := - if UniformContinuous f then pkg.denseInducing.extend f else fun x => f (pkg.dense.some x) + if UniformContinuous f then pkg.isDenseInducing.extend f else fun x => f (pkg.dense.some x) variable {f : α → β} -theorem extend_def (hf : UniformContinuous f) : pkg.extend f = pkg.denseInducing.extend f := +theorem extend_def (hf : UniformContinuous f) : pkg.extend f = pkg.isDenseInducing.extend f := if_pos hf theorem extend_coe [T2Space β] (hf : UniformContinuous f) (a : α) : (pkg.extend f) (ι a) = f a := by rw [pkg.extend_def hf] - exact pkg.denseInducing.extend_eq hf.continuous a + exact pkg.isDenseInducing.extend_eq hf.continuous a variable [CompleteSpace β] @@ -276,17 +276,17 @@ theorem compare_comp_eq_compare (γ : Type*) [TopologicalSpace γ] letI := pkg.uniformStruct.toTopologicalSpace letI := pkg'.uniformStruct.toTopologicalSpace (∀ a : pkg.space, - Filter.Tendsto f (Filter.comap pkg.coe (𝓝 a)) (𝓝 ((pkg.denseInducing.extend f) a))) → - pkg.denseInducing.extend f ∘ pkg'.compare pkg = pkg'.denseInducing.extend f := by + Filter.Tendsto f (Filter.comap pkg.coe (𝓝 a)) (𝓝 ((pkg.isDenseInducing.extend f) a))) → + pkg.isDenseInducing.extend f ∘ pkg'.compare pkg = pkg'.isDenseInducing.extend f := by let _ := pkg'.uniformStruct let _ := pkg.uniformStruct intro h - have (x : α) : (pkg.denseInducing.extend f ∘ pkg'.compare pkg) (pkg'.coe x) = f x := by - simp only [Function.comp_apply, compare_coe, DenseInducing.extend_eq _ cont_f, implies_true] - apply (DenseInducing.extend_unique (AbstractCompletion.denseInducing _) this + have (x : α) : (pkg.isDenseInducing.extend f ∘ pkg'.compare pkg) (pkg'.coe x) = f x := by + simp only [Function.comp_apply, compare_coe, IsDenseInducing.extend_eq _ cont_f, implies_true] + apply (IsDenseInducing.extend_unique (AbstractCompletion.isDenseInducing _) this (Continuous.comp _ (uniformContinuous_compare pkg' pkg).continuous )).symm - apply DenseInducing.continuous_extend - exact fun a ↦ ⟨(pkg.denseInducing.extend f) a, h a⟩ + apply IsDenseInducing.continuous_extend + exact fun a ↦ ⟨(pkg.isDenseInducing.extend f) a, h a⟩ end Compare diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 79943bfbcdbcf..4bfc074965c87 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -32,15 +32,15 @@ theorem UniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [ ClosedEmbedding f := ⟨hf.embedding, hf.toUniformInducing.isComplete_range.isClosed⟩ -namespace DenseInducing +namespace IsDenseInducing open Filter variable [TopologicalSpace α] {β : Type*} [TopologicalSpace β] variable {γ : Type*} [UniformSpace γ] [CompleteSpace γ] [T0Space γ] -theorem continuous_extend_of_cauchy {e : α → β} {f : α → γ} (de : DenseInducing e) +theorem continuous_extend_of_cauchy {e : α → β} {f : α → γ} (de : IsDenseInducing e) (h : ∀ b : β, Cauchy (map f (comap e <| 𝓝 b))) : Continuous (de.extend f) := de.continuous_extend fun b => CompleteSpace.complete (h b) -end DenseInducing +end IsDenseInducing diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 25ec9b9b1bb0f..b038f0744d751 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -180,15 +180,15 @@ theorem denseRange_pureCauchy : DenseRange (pureCauchy : α → CauchyFilter α) ⟨mem_range_self y, hy⟩ exact ⟨_, this⟩ -theorem denseInducing_pureCauchy : DenseInducing (pureCauchy : α → CauchyFilter α) := - uniformInducing_pureCauchy.denseInducing denseRange_pureCauchy +theorem isDenseInducing_pureCauchy : IsDenseInducing (pureCauchy : α → CauchyFilter α) := + uniformInducing_pureCauchy.isDenseInducing denseRange_pureCauchy theorem denseEmbedding_pureCauchy : DenseEmbedding (pureCauchy : α → CauchyFilter α) := uniformEmbedding_pureCauchy.denseEmbedding denseRange_pureCauchy theorem nonempty_cauchyFilter_iff : Nonempty (CauchyFilter α) ↔ Nonempty α := by constructor <;> rintro ⟨c⟩ - · have := eq_univ_iff_forall.1 denseEmbedding_pureCauchy.toDenseInducing.closure_range c + · have := eq_univ_iff_forall.1 denseEmbedding_pureCauchy.toIsDenseInducing.closure_range c obtain ⟨_, ⟨_, a, _⟩⟩ := mem_closure_iff.1 this _ isOpen_univ trivial exact ⟨a⟩ · exact ⟨pureCauchy c⟩ @@ -224,7 +224,7 @@ open Classical in /-- Extend a uniformly continuous function `α → β` to a function `CauchyFilter α → β`. Outputs junk when `f` is not uniformly continuous. -/ def extend (f : α → β) : CauchyFilter α → β := - if UniformContinuous f then denseInducing_pureCauchy.extend f + if UniformContinuous f then isDenseInducing_pureCauchy.extend f else fun x => f (nonempty_cauchyFilter_iff.1 ⟨x⟩).some section T0Space @@ -365,7 +365,7 @@ theorem coe_injective [T0Space α] : Function.Injective ((↑) : α → Completi variable {α} -theorem denseInducing_coe : DenseInducing ((↑) : α → Completion α) := +theorem isDenseInducing_coe : IsDenseInducing ((↑) : α → Completion α) := { (uniformInducing_coe α).inducing with dense := denseRange_coe } /-- The uniform bijection between a complete space and its uniform completion. -/ @@ -375,10 +375,10 @@ def UniformCompletion.completeEquivSelf [CompleteSpace α] [T0Space α] : Comple open TopologicalSpace instance separableSpace_completion [SeparableSpace α] : SeparableSpace (Completion α) := - Completion.denseInducing_coe.separableSpace + Completion.isDenseInducing_coe.separableSpace theorem denseEmbedding_coe [T0Space α] : DenseEmbedding ((↑) : α → Completion α) := - { denseInducing_coe with inj := separated_pureCauchy_injective } + { isDenseInducing_coe with inj := separated_pureCauchy_injective } theorem denseRange_coe₂ : DenseRange fun x : α × β => ((x.1 : Completion α), (x.2 : Completion β)) := diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index ed90ccfc3618a..b8882ff944db5 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -111,8 +111,8 @@ theorem UniformInducing.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [Uni UniformInducing fun p : α × β => (e₁ p.1, e₂ p.2) := ⟨by simp [Function.comp_def, uniformity_prod, ← h₁.1, ← h₂.1, comap_inf, comap_comap]⟩ -theorem UniformInducing.denseInducing {f : α → β} (h : UniformInducing f) (hd : DenseRange f) : - DenseInducing f := +theorem UniformInducing.isDenseInducing {f : α → β} (h : UniformInducing f) (hd : DenseRange f) : + IsDenseInducing f := { dense := hd induced := h.inducing.induced } @@ -232,7 +232,7 @@ theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopolo isClosed_range := isClosed_range_of_spaced_out hs hf } theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α → β} (b : β) - (he₁ : UniformInducing e) (he₂ : DenseInducing e) (hs : s ∈ 𝓤 α) : + (he₁ : UniformInducing e) (he₂ : IsDenseInducing e) (hs : s ∈ 𝓤 α) : ∃ a, closure (e '' { a' | (a, a') ∈ s }) ∈ 𝓝 b := by obtain ⟨U, ⟨hU, hUo, hsymm⟩, hs⟩ : ∃ U, (U ∈ 𝓤 β ∧ IsOpen U ∧ SymmetricRel U) ∧ Prod.map e e ⁻¹' U ⊆ s := by @@ -408,11 +408,11 @@ variable {α : Type*} {β : Type*} {γ : Type*} [UniformSpace α] [UniformSpace {e : β → α} (h_e : UniformInducing e) (h_dense : DenseRange e) {f : β → γ} (h_f : UniformContinuous f) -local notation "ψ" => DenseInducing.extend (UniformInducing.denseInducing h_e h_dense) f +local notation "ψ" => IsDenseInducing.extend (UniformInducing.isDenseInducing h_e h_dense) f include h_e h_dense h_f in theorem uniformly_extend_exists [CompleteSpace γ] (a : α) : ∃ c, Tendsto f (comap e (𝓝 a)) (𝓝 c) := - let de := h_e.denseInducing h_dense + let de := h_e.isDenseInducing h_dense have : Cauchy (𝓝 a) := cauchy_nhds have : Cauchy (comap e (𝓝 a)) := this.comap' (le_of_eq h_e.comap_uniformity) (de.comap_nhds_neBot _) @@ -437,7 +437,7 @@ theorem uniform_extend_subtype [CompleteSpace γ] {p : α → Prop} {e : α → include h_e h_f in theorem uniformly_extend_spec [CompleteSpace γ] (a : α) : Tendsto f (comap e (𝓝 a)) (𝓝 (ψ a)) := by - simpa only [DenseInducing.extend] using + simpa only [IsDenseInducing.extend] using tendsto_nhds_limUnder (uniformly_extend_exists h_e ‹_› h_f _) include h_f in @@ -446,7 +446,7 @@ theorem uniformContinuous_uniformly_extend [CompleteSpace γ] : UniformContinuou have h_pnt : ∀ {a m}, m ∈ 𝓝 a → ∃ c ∈ f '' (e ⁻¹' m), (c, ψ a) ∈ s ∧ (ψ a, c) ∈ s := fun {a m} hm => have nb : NeBot (map f (comap e (𝓝 a))) := - ((h_e.denseInducing h_dense).comap_nhds_neBot _).map _ + ((h_e.isDenseInducing h_dense).comap_nhds_neBot _).map _ have : f '' (e ⁻¹' m) ∩ ({ c | (c, ψ a) ∈ s } ∩ { c | (ψ a, c) ∈ s }) ∈ map f (comap e (𝓝 a)) := inter_mem (image_mem_map <| preimage_mem_comap <| hm) @@ -471,9 +471,9 @@ variable [T0Space γ] include h_f in theorem uniformly_extend_of_ind (b : β) : ψ (e b) = f b := - DenseInducing.extend_eq_at _ h_f.continuous.continuousAt + IsDenseInducing.extend_eq_at _ h_f.continuous.continuousAt theorem uniformly_extend_unique {g : α → γ} (hg : ∀ b, g (e b) = f b) (hc : Continuous g) : ψ = g := - DenseInducing.extend_unique _ hg hc + IsDenseInducing.extend_unique _ hg hc end UniformExtension diff --git a/docs/overview.yaml b/docs/overview.yaml index e1d21817d53b5..d020bc5328e44 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -209,7 +209,7 @@ Topology: cluster point: 'ClusterPt' Hausdorff space: 'T2Space' sequential space: 'SequentialSpace' - extension by continuity: 'DenseInducing.extend' + extension by continuity: 'IsDenseInducing.extend' compactness in terms of filters: 'IsCompact' compactness in terms of open covers (Borel-Lebesgue): 'isCompact_iff_finite_subcover' connectedness: 'ConnectedSpace' diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 5a3fa4e0c5ec5..e8be7c1b7ded7 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -1001,8 +1001,8 @@ DedekindDomain.ProdAdicCompletions.IsFiniteAdele.algebraMap' DenseEmbedding.mk' Dense.exists_ge' Dense.exists_le' -DenseInducing.extend_eq_at' -DenseInducing.mk' +IsDenseInducing.extend_eq_at' +IsDenseInducing.mk' Denumerable.lower_raise' Denumerable.raise_lower' deriv_add_const' From bc4aae561fbe8ce6a98cee62ac482dfb18fcac8d Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sun, 29 Sep 2024 00:04:21 +0000 Subject: [PATCH 028/174] chore: replace top by univ in closure_cycle_adjacent_swap (#17235) --- Mathlib/GroupTheory/Perm/Closure.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib/GroupTheory/Perm/Closure.lean b/Mathlib/GroupTheory/Perm/Closure.lean index 6fa1105deb178..9564bafaf6fcc 100644 --- a/Mathlib/GroupTheory/Perm/Closure.lean +++ b/Mathlib/GroupTheory/Perm/Closure.lean @@ -40,7 +40,7 @@ theorem closure_isCycle : closure { σ : Perm β | IsCycle σ } = ⊤ := by variable [DecidableEq α] [Fintype α] -theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.support = ⊤) (x : α) : +theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.support = univ) (x : α) : closure ({σ, swap x (σ x)} : Set (Perm α)) = ⊤ := by let H := closure ({σ, swap x (σ x)} : Set (Perm α)) have h3 : σ ∈ H := subset_closure (Set.mem_insert σ _) @@ -51,8 +51,7 @@ theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.su | zero => exact subset_closure (Set.mem_insert_of_mem _ (Set.mem_singleton _)) | succ n ih => convert H.mul_mem (H.mul_mem h3 ih) (H.inv_mem h3) - simp_rw [mul_swap_eq_swap_mul, mul_inv_cancel_right, pow_succ'] - rfl + simp_rw [mul_swap_eq_swap_mul, mul_inv_cancel_right, pow_succ', coe_mul, comp_apply] have step2 : ∀ n : ℕ, swap x ((σ ^ n) x) ∈ H := by intro n induction n with @@ -70,9 +69,9 @@ theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.su exact H.mul_mem (H.mul_mem (step1 n) ih) (step1 n) have step3 : ∀ y : α, swap x y ∈ H := by intro y - have hx : x ∈ (⊤ : Finset α) := Finset.mem_univ x + have hx : x ∈ univ := Finset.mem_univ x rw [← h2, mem_support] at hx - have hy : y ∈ (⊤ : Finset α) := Finset.mem_univ y + have hy : y ∈ univ := Finset.mem_univ y rw [← h2, mem_support] at hy cases' IsCycle.exists_pow_eq h1 hx hy with n hn rw [← hn] @@ -97,7 +96,7 @@ theorem closure_cycle_coprime_swap {n : ℕ} {σ : Perm α} (h0 : Nat.Coprime n closure ({σ, swap x ((σ ^ n) x)} : Set (Perm α)) = ⊤ := by rw [← Finset.card_univ, ← h2, ← h1.orderOf] at h0 cases' exists_pow_eq_self_of_coprime h0 with m hm - have h2' : (σ ^ n).support = ⊤ := Eq.trans (support_pow_coprime h0) h2 + have h2' : (σ ^ n).support = univ := Eq.trans (support_pow_coprime h0) h2 have h1' : IsCycle ((σ ^ n) ^ (m : ℤ)) := by rwa [← hm] at h1 replace h1' : IsCycle (σ ^ n) := h1'.of_pow (le_trans (support_pow_le σ n) (ge_of_eq (congr_arg support hm))) From 88304cedde34c73cab5c18a89b23663527d957b0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 29 Sep 2024 00:04:22 +0000 Subject: [PATCH 029/174] fix(Matrix/Invertible): deprecated aliases should remain `protected` (#17242) The originals were protected, and the deprecations break downstream code using the matching lemmas in the root namespace. A second patch to #16590. --- Mathlib/Data/Matrix/Invertible.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/Data/Matrix/Invertible.lean b/Mathlib/Data/Matrix/Invertible.lean index b68c2fb925d3d..d5581bd17d0de 100644 --- a/Mathlib/Data/Matrix/Invertible.lean +++ b/Mathlib/Data/Matrix/Invertible.lean @@ -48,12 +48,14 @@ protected theorem invOf_mul_cancel_right (A : Matrix m n α) (B : Matrix n n α) protected theorem mul_invOf_cancel_right (A : Matrix m n α) (B : Matrix n n α) [Invertible B] : A * B * ⅟ B = A := by rw [Matrix.mul_assoc, mul_invOf_self, Matrix.mul_one] -@[deprecated (since := "2024-09-07")] alias invOf_mul_self_assoc := Matrix.invOf_mul_cancel_left -@[deprecated (since := "2024-09-07")] alias mul_invOf_self_assoc := Matrix.mul_invOf_cancel_left @[deprecated (since := "2024-09-07")] -alias mul_invOf_mul_self_cancel := Matrix.invOf_mul_cancel_right +protected alias invOf_mul_self_assoc := Matrix.invOf_mul_cancel_left @[deprecated (since := "2024-09-07")] -alias mul_mul_invOf_self_cancel := Matrix.mul_invOf_cancel_right +protected alias mul_invOf_self_assoc := Matrix.mul_invOf_cancel_left +@[deprecated (since := "2024-09-07")] +protected alias mul_invOf_mul_self_cancel := Matrix.invOf_mul_cancel_right +@[deprecated (since := "2024-09-07")] +protected alias mul_mul_invOf_self_cancel := Matrix.mul_invOf_cancel_right section ConjTranspose variable [StarRing α] (A : Matrix n n α) From 197417dc8fc00c22a2e575192767564190f1417e Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 29 Sep 2024 00:39:20 +0000 Subject: [PATCH 030/174] chore(scripts): update nolints.json (#17245) I am happy to remove some nolints for you! --- scripts/nolints.json | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index c85078ae6db23..d897d5a3db174 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -285,7 +285,6 @@ ["docBlame", "PowerBasis.basis"], ["docBlame", "PowerBasis.dim"], ["docBlame", "PowerBasis.gen"], - ["docBlame", "PresheafOfModules.presheaf"], ["docBlame", "Pretrivialization.baseSet"], ["docBlame", "PrimeSpectrum.asIdeal"], ["docBlame", "ProbabilityTheory.«termEVar[_]»"], @@ -543,7 +542,6 @@ ["docBlame", "Order.PFilter.dual"], ["docBlame", "PProd.mk.injArrow"], ["docBlame", "PicardLindelof.FunSpace.toFun"], - ["docBlame", "PresheafOfModules.Hom.hom"], ["docBlame", "Prod.mk.injArrow"], ["docBlame", "QuaternionAlgebra.Basis.i"], ["docBlame", "QuaternionAlgebra.Basis.j"], @@ -755,4 +753,4 @@ ["docBlame", "Mathlib.Meta.NormNum.evalAdd.core.ratArm"], ["docBlame", "Mathlib.Meta.NormNum.evalMul.core.intArm"], ["docBlame", "Mathlib.Meta.NormNum.evalMul.core.ratArm"], - ["unusedArguments", "Combinator.K"]] \ No newline at end of file + ["unusedArguments", "Combinator.K"]] From 4cd4cfb8ebb1f63252cdeaf0f82009b4cf19ae06 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 29 Sep 2024 03:18:05 +0000 Subject: [PATCH 031/174] chore(UniformSpace): move `ofFun` to a new file (#16737) Also drop `@[mono]` on `hasSum_mono`. --- Mathlib.lean | 1 + .../Topology/Algebra/InfiniteSum/Order.lean | 2 +- .../Topology/ContinuousFunction/Bounded.lean | 1 + Mathlib/Topology/EMetricSpace/Defs.lean | 2 +- Mathlib/Topology/Support.lean | 1 + .../Topology/UniformSpace/AbsoluteValue.lean | 2 +- Mathlib/Topology/UniformSpace/Basic.lean | 29 +---------- Mathlib/Topology/UniformSpace/OfFun.lean | 52 +++++++++++++++++++ scripts/noshake.json | 3 +- 9 files changed, 61 insertions(+), 32 deletions(-) create mode 100644 Mathlib/Topology/UniformSpace/OfFun.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1acceb28d4138..1431bd8d7f883 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4801,6 +4801,7 @@ import Mathlib.Topology.UniformSpace.Completion import Mathlib.Topology.UniformSpace.Equicontinuity import Mathlib.Topology.UniformSpace.Equiv import Mathlib.Topology.UniformSpace.Matrix +import Mathlib.Topology.UniformSpace.OfFun import Mathlib.Topology.UniformSpace.Pi import Mathlib.Topology.UniformSpace.Separation import Mathlib.Topology.UniformSpace.UniformConvergence diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index 53773dace6081..88fad0d581136 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -40,7 +40,7 @@ variable [OrderedCommMonoid α] [TopologicalSpace α] [OrderClosedTopology α] { theorem hasProd_le (h : ∀ i, f i ≤ g i) (hf : HasProd f a₁) (hg : HasProd g a₂) : a₁ ≤ a₂ := le_of_tendsto_of_tendsto' hf hg fun _ ↦ prod_le_prod' fun i _ ↦ h i -@[to_additive (attr := mono)] +@[to_additive] theorem hasProd_mono (hf : HasProd f a₁) (hg : HasProd g a₂) (h : f ≤ g) : a₁ ≤ a₂ := hasProd_le h hf hg diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index 619fbc2c6b8a5..9b505b1746655 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -9,6 +9,7 @@ import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Topology.Bornology.BoundedOperation +import Mathlib.Tactic.Monotonicity /-! # Bounded continuous functions diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 5afe753e4baa3..12ff875ac5f87 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -4,7 +4,7 @@ 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.Inv -import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Topology.UniformSpace.OfFun /-! # Extended metric spaces diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 8f8bcb0d14606..e5e52b3d9982f 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. 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 diff --git a/Mathlib/Topology/UniformSpace/AbsoluteValue.lean b/Mathlib/Topology/UniformSpace/AbsoluteValue.lean index fb21530f3f80a..55efcadaf7133 100644 --- a/Mathlib/Topology/UniformSpace/AbsoluteValue.lean +++ b/Mathlib/Topology/UniformSpace/AbsoluteValue.lean @@ -5,7 +5,7 @@ Authors: Patrick Massot -/ import Mathlib.Algebra.Order.AbsoluteValue import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Topology.UniformSpace.OfFun /-! # Uniform structure induced by an absolute value diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index eca83bcfea729..b57b69003857a 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Order.Filter.SmallSets -import Mathlib.Tactic.Monotonicity +import Mathlib.Tactic.Monotonicity.Basic import Mathlib.Topology.Compactness.Compact import Mathlib.Topology.NhdsSet import Mathlib.Algebra.Group.Defs @@ -367,33 +367,6 @@ theorem UniformSpace.replaceTopology_eq {α : Type*} [i : TopologicalSpace α] ( (h : i = u.toTopologicalSpace) : u.replaceTopology h = u := UniformSpace.ext rfl --- Porting note: rfc: use `UniformSpace.Core.mkOfBasis`? This will change defeq here and there -/-- Define a `UniformSpace` using a "distance" function. The function can be, e.g., the -distance in a (usual or extended) metric space or an absolute value on a ring. -/ -def UniformSpace.ofFun {α : Type u} {β : Type v} [OrderedAddCommMonoid β] - (d : α → α → β) (refl : ∀ x, d x x = 0) (symm : ∀ x y, d x y = d y x) - (triangle : ∀ x y z, d x z ≤ d x y + d y z) - (half : ∀ ε > (0 : β), ∃ δ > (0 : β), ∀ x < δ, ∀ y < δ, x + y < ε) : - UniformSpace α := - .ofCore - { uniformity := ⨅ r > 0, 𝓟 { x | d x.1 x.2 < r } - refl := le_iInf₂ fun r hr => principal_mono.2 <| idRel_subset.2 fun x => by simpa [refl] - symm := tendsto_iInf_iInf fun r => tendsto_iInf_iInf fun _ => tendsto_principal_principal.2 - fun x hx => by rwa [mem_setOf, symm] - comp := le_iInf₂ fun r hr => let ⟨δ, h0, hδr⟩ := half r hr; le_principal_iff.2 <| - mem_of_superset - (mem_lift' <| mem_iInf_of_mem δ <| mem_iInf_of_mem h0 <| mem_principal_self _) - fun (x, z) ⟨y, h₁, h₂⟩ => (triangle _ _ _).trans_lt (hδr _ h₁ _ h₂) } - -theorem UniformSpace.hasBasis_ofFun {α : Type u} {β : Type v} [LinearOrderedAddCommMonoid β] - (h₀ : ∃ x : β, 0 < x) (d : α → α → β) (refl : ∀ x, d x x = 0) (symm : ∀ x y, d x y = d y x) - (triangle : ∀ x y z, d x z ≤ d x y + d y z) - (half : ∀ ε > (0 : β), ∃ δ > (0 : β), ∀ x < δ, ∀ y < δ, x + y < ε) : - 𝓤[.ofFun d refl symm triangle half].HasBasis ((0 : β) < ·) (fun ε => { x | d x.1 x.2 < ε }) := - hasBasis_biInf_principal' - (fun ε₁ h₁ ε₂ h₂ => ⟨min ε₁ ε₂, lt_min h₁ h₂, fun _x hx => lt_of_lt_of_le hx (min_le_left _ _), - fun _x hx => lt_of_lt_of_le hx (min_le_right _ _)⟩) h₀ - section UniformSpace variable [UniformSpace α] diff --git a/Mathlib/Topology/UniformSpace/OfFun.lean b/Mathlib/Topology/UniformSpace/OfFun.lean new file mode 100644 index 0000000000000..6a1ff605f190e --- /dev/null +++ b/Mathlib/Topology/UniformSpace/OfFun.lean @@ -0,0 +1,52 @@ +/- +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.UniformSpace.Basic +import Mathlib.Algebra.Order.Monoid.Defs + +/-! +# Construct a `UniformSpace` from a `dist`-like function + +In this file we provide a constructor for `UniformSpace` +given a `dist`-like function + +## TODO + +RFC: use `UniformSpace.Core.mkOfBasis`? This will change defeq here and there +-/ + +open Filter Set +open scoped Uniformity + +variable {X M : Type*} + +namespace UniformSpace + +/-- Define a `UniformSpace` using a "distance" function. The function can be, e.g., the +distance in a (usual or extended) metric space or an absolute value on a ring. -/ +def ofFun [OrderedAddCommMonoid M] (d : X → X → M) (refl : ∀ x, d x x = 0) + (symm : ∀ x y, d x y = d y x) (triangle : ∀ x y z, d x z ≤ d x y + d y z) + (half : ∀ ε > (0 : M), ∃ δ > (0 : M), ∀ x < δ, ∀ y < δ, x + y < ε) : + UniformSpace X := + .ofCore + { uniformity := ⨅ r > 0, 𝓟 { x | d x.1 x.2 < r } + refl := le_iInf₂ fun r hr => principal_mono.2 <| idRel_subset.2 fun x => by simpa [refl] + symm := tendsto_iInf_iInf fun r => tendsto_iInf_iInf fun _ => tendsto_principal_principal.2 + fun x hx => by rwa [mem_setOf, symm] + comp := le_iInf₂ fun r hr => let ⟨δ, h0, hδr⟩ := half r hr; le_principal_iff.2 <| + mem_of_superset + (mem_lift' <| mem_iInf_of_mem δ <| mem_iInf_of_mem h0 <| mem_principal_self _) + fun (x, z) ⟨y, h₁, h₂⟩ => (triangle _ _ _).trans_lt (hδr _ h₁ _ h₂) } + +theorem hasBasis_ofFun [LinearOrderedAddCommMonoid M] + (h₀ : ∃ x : M, 0 < x) (d : X → X → M) (refl : ∀ x, d x x = 0) (symm : ∀ x y, d x y = d y x) + (triangle : ∀ x y z, d x z ≤ d x y + d y z) + (half : ∀ ε > (0 : M), ∃ δ > (0 : M), ∀ x < δ, ∀ y < δ, x + y < ε) : + 𝓤[.ofFun d refl symm triangle half].HasBasis ((0 : M) < ·) (fun ε => { x | d x.1 x.2 < ε }) := + hasBasis_biInf_principal' + (fun ε₁ h₁ ε₂ h₂ => ⟨min ε₁ ε₂, lt_min h₁ h₂, fun _x hx => lt_of_lt_of_le hx (min_le_left _ _), + fun _x hx => lt_of_lt_of_le hx (min_le_right _ _)⟩) h₀ + +end UniformSpace diff --git a/scripts/noshake.json b/scripts/noshake.json index 5335397efa8fc..735d7aacd17cf 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -194,7 +194,8 @@ "ignoreAll": ["Batteries.Tactic.Basic", "Mathlib.Mathport.Syntax", "Mathlib.Tactic.Linter"], "ignore": - {"Mathlib.Topology.Sheaves.Forget": ["Mathlib.Algebra.Category.Ring.Limits"], + {"Mathlib.Topology.UniformSpace.Basic": ["Mathlib.Tactic.Monotonicity.Basic"], + "Mathlib.Topology.Sheaves.Forget": ["Mathlib.Algebra.Category.Ring.Limits"], "Mathlib.Topology.Order.LeftRightNhds": ["Mathlib.Algebra.Ring.Pointwise.Set"], "Mathlib.Topology.Germ": ["Mathlib.Analysis.Normed.Module.Basic"], From d244ef2651a169162e866e63d8c0d5e4358746f2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 29 Sep 2024 03:27:12 +0000 Subject: [PATCH 032/174] fix: call checkSystem in `linarith` and `ring` (#16666) Without this change, linarith calls on large states can substantially overrun On the web editor, the test fails with ``` Tactic took much more than 750ms (1272ms) Tactic took much more than 1000ms (1267ms) Tactic took much more than 2250ms (2679ms) ``` Co-authored-by: Laurent Sartran Co-authored-by: Eric Wieser --- .../Linarith/Oracle/FourierMotzkin.lean | 22 +- .../Oracle/SimplexAlgorithm/Gauss.lean | 5 +- .../SimplexAlgorithm/PositiveVector.lean | 4 +- .../SimplexAlgorithm/SimplexAlgorithm.lean | 5 +- Mathlib/Tactic/Linarith/Verification.lean | 1 + Mathlib/Tactic/Ring/Basic.lean | 216 ++++++++++-------- test/tactic_timeout.lean | 98 ++++++++ 7 files changed, 239 insertions(+), 112 deletions(-) create mode 100644 test/tactic_timeout.lean diff --git a/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean b/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean index 053e563cbbe5e..d1b3479c6c896 100644 --- a/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean +++ b/Mathlib/Tactic/Linarith/Oracle/FourierMotzkin.lean @@ -259,7 +259,7 @@ The linarith monad extends an exceptional monad with a `LinarithData` state. An exception produces a contradictory `PComp`. -/ abbrev LinarithM : Type → Type := - StateT LinarithData (ExceptT PComp Id) + StateT LinarithData (ExceptT PComp Lean.Core.CoreM) /-- Returns the current max variable. -/ def getMaxVar : LinarithM ℕ := @@ -273,7 +273,7 @@ def getPCompSet : LinarithM PCompSet := def validate : LinarithM Unit := do match (← getPCompSet).toList.find? (fun p : PComp => p.isContr) with | none => return () - | some c => throw c + | some c => throwThe _ c /-- Updates the current state with a new max variable and comparisons, @@ -305,9 +305,12 @@ from the `linarith` state. -/ def elimVarM (a : ℕ) : LinarithM Unit := do let vs ← getMaxVar - if (a ≤ vs) then (do + if (a ≤ vs) then + Lean.Core.checkSystem decl_name%.toString let ⟨pos, neg, notPresent⟩ := splitSetByVarSign a (← getPCompSet) - update (vs - 1) (pos.foldl (fun s p => s.union (elimWithSet a p neg)) notPresent)) + update (vs - 1) (← pos.foldlM (fun s p => do + Lean.Core.checkSystem decl_name%.toString + pure (s.union (elimWithSet a p neg))) notPresent) else pure () @@ -328,9 +331,12 @@ def mkLinarithData (hyps : List Comp) (maxVar : ℕ) : LinarithData := /-- An oracle that uses Fourier-Motzkin elimination. -/ def CertificateOracle.fourierMotzkin : CertificateOracle where - produceCertificate hyps maxVar := match ExceptT.run - (StateT.run (do validate; elimAllVarsM : LinarithM Unit) (mkLinarithData hyps maxVar)) with - | (Except.ok _) => failure - | (Except.error contr) => return contr.src.flatten + produceCertificate hyps maxVar := do + let linarithData := mkLinarithData hyps maxVar + let result ← + (ExceptT.run (StateT.run (do validate; elimAllVarsM : LinarithM Unit) linarithData) : _) + match result with + | (Except.ok _) => failure + | (Except.error contr) => return contr.src.flatten end Linarith diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Gauss.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Gauss.lean index ab46c5726b44f..93cad0a43aa77 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Gauss.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Gauss.lean @@ -15,7 +15,7 @@ solution which is done by standard Gaussian Elimination algorithm implemented in namespace Linarith.SimplexAlgorithm.Gauss /-- The monad for the Gaussian Elimination algorithm. -/ -abbrev GaussM (n m : Nat) (matType : Nat → Nat → Type) := StateM <| matType n m +abbrev GaussM (n m : Nat) (matType : Nat → Nat → Type) := StateT (matType n m) Lean.CoreM variable {n m : Nat} {matType : Nat → Nat → Type} [UsableInSimplexAlgorithm matType] @@ -35,6 +35,7 @@ def getTableauImp : GaussM n m matType <| Tableau matType := do let mut col : Nat := 0 while row < n && col < m do + Lean.Core.checkSystem decl_name%.toString match ← findNonzeroRow row col with | .none => free := free.push col @@ -74,7 +75,7 @@ Given matrix `A`, solves the linear equation `A x = 0` and returns the solution some variables are free and others (basic) variable are expressed as linear combinations of the free ones. -/ -def getTableau (A : matType n m) : Tableau matType := Id.run do +def getTableau (A : matType n m) : Lean.CoreM (Tableau matType) := do return (← getTableauImp.run A).fst end Linarith.SimplexAlgorithm.Gauss diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean index c93e2e33f9b24..97f8d6d475622 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean @@ -90,10 +90,10 @@ def findPositiveVector {n m : Nat} {matType : Nat → Nat → Type} [UsableInSim /- Using Gaussian elimination split variable into free and basic forming the tableau that will be operated by the Simplex Algorithm. -/ - let initTableau := Gauss.getTableau B + let initTableau ← Gauss.getTableau B /- Run the Simplex Algorithm and extract the solution. -/ - let res := runSimplexAlgorithm.run initTableau + let res ← runSimplexAlgorithm.run initTableau if res.fst.isOk then return extractSolution res.snd else diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean index 924814fa69fce..9858e94c23f14 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean @@ -21,7 +21,7 @@ inductive SimplexAlgorithmException /-- The monad for the Simplex Algorithm. -/ abbrev SimplexAlgorithmM (matType : Nat → Nat → Type) [UsableInSimplexAlgorithm matType] := - ExceptT SimplexAlgorithmException <| StateM (Tableau matType) + ExceptT SimplexAlgorithmException <| StateT (Tableau matType) Lean.CoreM variable {matType : Nat → Nat → Type} [UsableInSimplexAlgorithm matType] @@ -77,7 +77,7 @@ def chooseEnteringVar : SimplexAlgorithmM matType Nat := do /- If there is no such variable the solution does not exist for sure. -/ match enterIdxOpt with - | .none => throw SimplexAlgorithmException.infeasible + | .none => throwThe SimplexAlgorithmException SimplexAlgorithmException.infeasible | .some enterIdx => return enterIdx /-- @@ -116,6 +116,7 @@ such exists. -/ def runSimplexAlgorithm : SimplexAlgorithmM matType Unit := do while !(← checkSuccess) do + Lean.Core.checkSystem decl_name%.toString let ⟨exitIdx, enterIdx⟩ ← choosePivots doPivotOperation exitIdx enterIdx diff --git a/Mathlib/Tactic/Linarith/Verification.lean b/Mathlib/Tactic/Linarith/Verification.lean index 9cda494376353..b8702e24fc4f5 100644 --- a/Mathlib/Tactic/Linarith/Verification.lean +++ b/Mathlib/Tactic/Linarith/Verification.lean @@ -191,6 +191,7 @@ def proveFalseByLinarith (transparency : TransparencyMode) (oracle : Certificate | _, [] => throwError "no args to linarith" | g, l@(h::_) => do trace[linarith.detail] "Beginning work in `proveFalseByLinarith`." + Lean.Core.checkSystem decl_name%.toString -- for the elimination to work properly, we must add a proof of `-1 < 0` to the list, -- along with negated equality proofs. let l' ← addNegEqProofs l diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 459dfa19e288a..498f20124180d 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -307,6 +307,11 @@ theorem add_overlap_pf_zero (x : R) (e) : IsNat (a + b) (nat_lit 0) → IsNat (x ^ e * a + x ^ e * b) (nat_lit 0) | ⟨h⟩ => ⟨by simp [h, ← mul_add]⟩ +-- TODO: decide if this is a good idea globally in +-- https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/.60MonadLift.20Option.20.28OptionT.20m.29.60/near/469097834 +private local instance {m} [Pure m] : MonadLift Option (OptionT m) where + monadLift f := .mk <| pure f + /-- Given monomials `va, vb`, attempts to add them together to get another monomial. If the monomials are not compatible, returns `none`. @@ -314,7 +319,8 @@ For example, `xy + 2xy = 3xy` is a `.nonzero` overlap, while `xy + xz` returns ` and `xy + -xy = 0` is a `.zero` overlap. -/ def evalAddOverlap {a b : Q($α)} (va : ExProd sα a) (vb : ExProd sα b) : - Option (Overlap sα q($a + $b)) := + OptionT Lean.Core.CoreM (Overlap sα q($a + $b)) := do + Lean.Core.checkSystem decl_name%.toString match va, vb with | .const za ha, .const zb hb => do let ra := Result.ofRawRat za a ha; let rb := Result.ofRawRat zb b hb @@ -331,7 +337,7 @@ def evalAddOverlap {a b : Q($α)} (va : ExProd sα a) (vb : ExProd sα b) : | .zero p => pure <| .zero (q(add_overlap_pf_zero $a₁ $a₂ $p) : Expr) | .nonzero ⟨_, vc, p⟩ => pure <| .nonzero ⟨_, .mul va₁ va₂ vc, (q(add_overlap_pf $a₁ $a₂ $p) : Expr)⟩ - | _, _ => none + | _, _ => OptionT.fail theorem add_pf_zero_add (b : R) : 0 + b = b := by simp @@ -359,25 +365,26 @@ theorem add_pf_add_gt (b₁ : R) (_ : a + b₂ = c) : a + (b₁ + b₂) = b₁ + * `(a₁ + a₂) + (b₁ + b₂) = b₁ + ((a₁ + a₂) + b₂)` (if not `a₁.lt b₁`) -/ partial def evalAdd {a b : Q($α)} (va : ExSum sα a) (vb : ExSum sα b) : - Result (ExSum sα) q($a + $b) := + Lean.Core.CoreM <| Result (ExSum sα) q($a + $b) := do + Lean.Core.checkSystem decl_name%.toString match va, vb with - | .zero, vb => ⟨b, vb, q(add_pf_zero_add $b)⟩ - | va, .zero => ⟨a, va, q(add_pf_add_zero $a)⟩ + | .zero, vb => return ⟨b, vb, q(add_pf_zero_add $b)⟩ + | va, .zero => return ⟨a, va, q(add_pf_add_zero $a)⟩ | .add (a := a₁) (b := _a₂) va₁ va₂, .add (a := b₁) (b := _b₂) vb₁ vb₂ => - match evalAddOverlap sα va₁ vb₁ with + match ← (evalAddOverlap sα va₁ vb₁).run with | some (.nonzero ⟨_, vc₁, pc₁⟩) => - let ⟨_, vc₂, pc₂⟩ := evalAdd va₂ vb₂ - ⟨_, .add vc₁ vc₂, q(add_pf_add_overlap $pc₁ $pc₂)⟩ + let ⟨_, vc₂, pc₂⟩ ← evalAdd va₂ vb₂ + return ⟨_, .add vc₁ vc₂, q(add_pf_add_overlap $pc₁ $pc₂)⟩ | some (.zero pc₁) => - let ⟨c₂, vc₂, pc₂⟩ := evalAdd va₂ vb₂ - ⟨c₂, vc₂, q(add_pf_add_overlap_zero $pc₁ $pc₂)⟩ + let ⟨c₂, vc₂, pc₂⟩ ← evalAdd va₂ vb₂ + return ⟨c₂, vc₂, q(add_pf_add_overlap_zero $pc₁ $pc₂)⟩ | none => if let .lt := va₁.cmp vb₁ then - let ⟨_c, vc, (pc : Q($_a₂ + ($b₁ + $_b₂) = $_c))⟩ := evalAdd va₂ vb - ⟨_, .add va₁ vc, q(add_pf_add_lt $a₁ $pc)⟩ + let ⟨_c, vc, (pc : Q($_a₂ + ($b₁ + $_b₂) = $_c))⟩ ← evalAdd va₂ vb + return ⟨_, .add va₁ vc, q(add_pf_add_lt $a₁ $pc)⟩ else - let ⟨_c, vc, (pc : Q($a₁ + $_a₂ + $_b₂ = $_c))⟩ := evalAdd va vb₂ - ⟨_, .add vb₁ vc, q(add_pf_add_gt $b₁ $pc)⟩ + let ⟨_c, vc, (pc : Q($a₁ + $_a₂ + $_b₂ = $_c))⟩ ← evalAdd va vb₂ + return ⟨_, .add vb₁ vc, q(add_pf_add_gt $b₁ $pc)⟩ theorem one_mul (a : R) : (nat_lit 1).rawCast * a = a := by simp [Nat.rawCast] @@ -406,37 +413,38 @@ theorem mul_pp_pf_overlap {ea eb e : ℕ} (x : R) (_ : ea + eb = e) (_ : a₂ * * `(a₁ * a₂) * (b₁ * b₂) = b₁ * ((a₁ * a₂) * b₂)` (if not `a₁.lt b₁`) -/ partial def evalMulProd {a b : Q($α)} (va : ExProd sα a) (vb : ExProd sα b) : - Result (ExProd sα) q($a * $b) := + Lean.Core.CoreM <| Result (ExProd sα) q($a * $b) := do + Lean.Core.checkSystem decl_name%.toString match va, vb with | .const za ha, .const zb hb => if za = 1 then - ⟨b, .const zb hb, (q(one_mul $b) : Expr)⟩ + return ⟨b, .const zb hb, (q(one_mul $b) : Expr)⟩ else if zb = 1 then - ⟨a, .const za ha, (q(mul_one $a) : Expr)⟩ + return ⟨a, .const za ha, (q(mul_one $a) : Expr)⟩ else let ra := Result.ofRawRat za a ha; let rb := Result.ofRawRat zb b hb let rc := (NormNum.evalMul.core q($a * $b) q(HMul.hMul) _ _ q(CommSemiring.toSemiring) ra rb).get! let ⟨zc, hc⟩ := rc.toRatNZ.get! let ⟨c, pc⟩ := rc.toRawEq - ⟨c, .const zc hc, pc⟩ + return ⟨c, .const zc hc, pc⟩ | .mul (x := a₁) (e := a₂) va₁ va₂ va₃, .const _ _ => - let ⟨_, vc, pc⟩ := evalMulProd va₃ vb - ⟨_, .mul va₁ va₂ vc, (q(mul_pf_left $a₁ $a₂ $pc) : Expr)⟩ + let ⟨_, vc, pc⟩ ← evalMulProd va₃ vb + return ⟨_, .mul va₁ va₂ vc, (q(mul_pf_left $a₁ $a₂ $pc) : Expr)⟩ | .const _ _, .mul (x := b₁) (e := b₂) vb₁ vb₂ vb₃ => - let ⟨_, vc, pc⟩ := evalMulProd va vb₃ - ⟨_, .mul vb₁ vb₂ vc, (q(mul_pf_right $b₁ $b₂ $pc) : Expr)⟩ - | .mul (x := xa) (e := ea) vxa vea va₂, .mul (x := xb) (e := eb) vxb veb vb₂ => Id.run do + let ⟨_, vc, pc⟩ ← evalMulProd va vb₃ + return ⟨_, .mul vb₁ vb₂ vc, (q(mul_pf_right $b₁ $b₂ $pc) : Expr)⟩ + | .mul (x := xa) (e := ea) vxa vea va₂, .mul (x := xb) (e := eb) vxb veb vb₂ => do if vxa.eq vxb then - if let some (.nonzero ⟨_, ve, pe⟩) := evalAddOverlap sℕ vea veb then - let ⟨_, vc, pc⟩ := evalMulProd va₂ vb₂ + if let some (.nonzero ⟨_, ve, pe⟩) ← (evalAddOverlap sℕ vea veb).run then + let ⟨_, vc, pc⟩ ← evalMulProd va₂ vb₂ return ⟨_, .mul vxa ve vc, (q(mul_pp_pf_overlap $xa $pe $pc) : Expr)⟩ if let .lt := (vxa.cmp vxb).then (vea.cmp veb) then - let ⟨_, vc, pc⟩ := evalMulProd va₂ vb - ⟨_, .mul vxa vea vc, (q(mul_pf_left $xa $ea $pc) : Expr)⟩ + let ⟨_, vc, pc⟩ ← evalMulProd va₂ vb + return ⟨_, .mul vxa vea vc, (q(mul_pf_left $xa $ea $pc) : Expr)⟩ else - let ⟨_, vc, pc⟩ := evalMulProd va vb₂ - ⟨_, .mul vxb veb vc, (q(mul_pf_right $xb $eb $pc) : Expr)⟩ + let ⟨_, vc, pc⟩ ← evalMulProd va vb₂ + return ⟨_, .mul vxb veb vc, (q(mul_pf_right $xb $eb $pc) : Expr)⟩ theorem mul_zero (a : R) : a * 0 = 0 := by simp @@ -449,14 +457,15 @@ theorem mul_add {d : R} (_ : (a : R) * b₁ = c₁) (_ : a * b₂ = c₂) (_ : c * `a * 0 = 0` * `a * (b₁ + b₂) = (a * b₁) + (a * b₂)` -/ -def evalMul₁ {a b : Q($α)} (va : ExProd sα a) (vb : ExSum sα b) : Result (ExSum sα) q($a * $b) := +def evalMul₁ {a b : Q($α)} (va : ExProd sα a) (vb : ExSum sα b) : + Lean.Core.CoreM <| Result (ExSum sα) q($a * $b) := do match vb with - | .zero => ⟨_, .zero, q(mul_zero $a)⟩ + | .zero => return ⟨_, .zero, q(mul_zero $a)⟩ | .add vb₁ vb₂ => - let ⟨_, vc₁, pc₁⟩ := evalMulProd sα va vb₁ - let ⟨_, vc₂, pc₂⟩ := evalMul₁ va vb₂ - let ⟨_, vd, pd⟩ := evalAdd sα vc₁.toSum vc₂ - ⟨_, vd, q(mul_add $pc₁ $pc₂ $pd)⟩ + let ⟨_, vc₁, pc₁⟩ ← evalMulProd sα va vb₁ + let ⟨_, vc₂, pc₂⟩ ← evalMul₁ va vb₂ + let ⟨_, vd, pd⟩ ← evalAdd sα vc₁.toSum vc₂ + return ⟨_, vd, q(mul_add $pc₁ $pc₂ $pd)⟩ theorem zero_mul (b : R) : 0 * b = 0 := by simp @@ -468,14 +477,15 @@ theorem add_mul {d : R} (_ : (a₁ : R) * b = c₁) (_ : a₂ * b = c₂) (_ : c * `0 * b = 0` * `(a₁ + a₂) * b = (a₁ * b) + (a₂ * b)` -/ -def evalMul {a b : Q($α)} (va : ExSum sα a) (vb : ExSum sα b) : Result (ExSum sα) q($a * $b) := +def evalMul {a b : Q($α)} (va : ExSum sα a) (vb : ExSum sα b) : + Lean.Core.CoreM <| Result (ExSum sα) q($a * $b) := do match va with - | .zero => ⟨_, .zero, q(zero_mul $b)⟩ + | .zero => return ⟨_, .zero, q(zero_mul $b)⟩ | .add va₁ va₂ => - let ⟨_, vc₁, pc₁⟩ := evalMul₁ sα va₁ vb - let ⟨_, vc₂, pc₂⟩ := evalMul va₂ vb - let ⟨_, vd, pd⟩ := evalAdd sα vc₁ vc₂ - ⟨_, vd, q(add_mul $pc₁ $pc₂ $pd)⟩ + let ⟨_, vc₁, pc₁⟩ ← evalMul₁ sα va₁ vb + let ⟨_, vc₂, pc₂⟩ ← evalMul va₂ vb + let ⟨_, vd, pd⟩ ← evalAdd sα vc₁ vc₂ + return ⟨_, vd, q(add_mul $pc₁ $pc₂ $pd)⟩ theorem natCast_nat (n) : ((Nat.rawCast n : ℕ) : R) = Nat.rawCast n := by simp @@ -552,11 +562,11 @@ def evalNSMul {a : Q(ℕ)} {b : Q($α)} (va : ExSum sℕ a) (vb : ExSum sα b) : if ← isDefEq sα sℕ then let ⟨_, va'⟩ := va.cast have _b : Q(ℕ) := b - let ⟨(_c : Q(ℕ)), vc, (pc : Q($a * $_b = $_c))⟩ := evalMul sα va' vb + let ⟨(_c : Q(ℕ)), vc, (pc : Q($a * $_b = $_c))⟩ ← evalMul sα va' vb pure ⟨_, vc, (q(smul_nat $pc) : Expr)⟩ else let ⟨_, va', pa'⟩ ← va.evalNatCast sα - let ⟨_, vc, pc⟩ := evalMul sα va' vb + let ⟨_, vc, pc⟩ ← evalMul sα va' vb pure ⟨_, vc, (q(smul_eq_cast $pa' $pc) : Expr)⟩ theorem neg_one_mul {R} [Ring R] {a b : R} (_ : (Int.negOfNat (nat_lit 1)).rawCast * a = b) : @@ -570,7 +580,9 @@ theorem neg_mul {R} [Ring R] (a₁ : R) (a₂) {a₃ b : R} * `-c = (-c)` (for `c` coefficient) * `-(a₁ * a₂) = a₁ * -a₂` -/ -def evalNegProd {a : Q($α)} (rα : Q(Ring $α)) (va : ExProd sα a) : Result (ExProd sα) q(-$a) := +def evalNegProd {a : Q($α)} (rα : Q(Ring $α)) (va : ExProd sα a) : + Lean.Core.CoreM <| Result (ExProd sα) q(-$a) := do + Lean.Core.checkSystem decl_name%.toString match va with | .const za ha => let lit : Q(ℕ) := mkRawNatLit 1 @@ -581,10 +593,10 @@ def evalNegProd {a : Q($α)} (rα : Q(Ring $α)) (va : ExProd sα a) : Result (E q(CommSemiring.toSemiring) rm ra).get! let ⟨zb, hb⟩ := rb.toRatNZ.get! let ⟨b, (pb : Q((Int.negOfNat (nat_lit 1)).rawCast * $a = $b))⟩ := rb.toRawEq - ⟨b, .const zb hb, (q(neg_one_mul (R := $α) $pb) : Expr)⟩ + return ⟨b, .const zb hb, (q(neg_one_mul (R := $α) $pb) : Expr)⟩ | .mul (x := a₁) (e := a₂) va₁ va₂ va₃ => - let ⟨_, vb, pb⟩ := evalNegProd rα va₃ - ⟨_, .mul va₁ va₂ vb, (q(neg_mul $a₁ $a₂ $pb) : Expr)⟩ + let ⟨_, vb, pb⟩ ← evalNegProd rα va₃ + return ⟨_, .mul va₁ va₂ vb, (q(neg_mul $a₁ $a₂ $pb) : Expr)⟩ theorem neg_zero {R} [Ring R] : -(0 : R) = 0 := by simp @@ -597,13 +609,14 @@ theorem neg_add {R} [Ring R] {a₁ a₂ b₁ b₂ : R} * `-0 = 0` (for `c` coefficient) * `-(a₁ + a₂) = -a₁ + -a₂` -/ -def evalNeg {a : Q($α)} (rα : Q(Ring $α)) (va : ExSum sα a) : Result (ExSum sα) q(-$a) := +def evalNeg {a : Q($α)} (rα : Q(Ring $α)) (va : ExSum sα a) : + Lean.Core.CoreM <| Result (ExSum sα) q(-$a) := do match va with - | .zero => ⟨_, .zero, (q(neg_zero (R := $α)) : Expr)⟩ + | .zero => return ⟨_, .zero, (q(neg_zero (R := $α)) : Expr)⟩ | .add va₁ va₂ => - let ⟨_, vb₁, pb₁⟩ := evalNegProd sα rα va₁ - let ⟨_, vb₂, pb₂⟩ := evalNeg rα va₂ - ⟨_, .add vb₁ vb₂, (q(neg_add $pb₁ $pb₂) : Expr)⟩ + let ⟨_, vb₁, pb₁⟩ ← evalNegProd sα rα va₁ + let ⟨_, vb₂, pb₂⟩ ← evalNeg rα va₂ + return ⟨_, .add vb₁ vb₂, (q(neg_add $pb₁ $pb₂) : Expr)⟩ theorem sub_pf {R} [Ring R] {a b c d : R} (_ : -b = c) (_ : a + c = d) : a - b = d := by subst_vars; simp [sub_eq_add_neg] @@ -613,10 +626,11 @@ theorem sub_pf {R} [Ring R] {a b c d : R} * `a - b = a + -b` -/ def evalSub {α : Q(Type u)} (sα : Q(CommSemiring $α)) {a b : Q($α)} - (rα : Q(Ring $α)) (va : ExSum sα a) (vb : ExSum sα b) : Result (ExSum sα) q($a - $b) := - let ⟨_c, vc, pc⟩ := evalNeg sα rα vb - let ⟨d, vd, (pd : Q($a + $_c = $d))⟩ := evalAdd sα va vc - ⟨d, vd, (q(sub_pf $pc $pd) : Expr)⟩ + (rα : Q(Ring $α)) (va : ExSum sα a) (vb : ExSum sα b) : + Lean.Core.CoreM <| Result (ExSum sα) q($a - $b) := do + let ⟨_c, vc, pc⟩ ← evalNeg sα rα vb + let ⟨d, vd, (pd : Q($a + $_c = $d))⟩ ← evalAdd sα va vc + return ⟨d, vd, (q(sub_pf $pc $pd) : Expr)⟩ theorem pow_prod_atom (a : R) (b) : a ^ b = (a + 0) ^ b * (nat_lit 1).rawCast := by simp @@ -718,22 +732,23 @@ into a sum of monomials. * `x ^ (2*n) = x ^ n * x ^ n` * `x ^ (2*n+1) = x ^ n * x ^ n * x` -/ -partial def evalPowNat {a : Q($α)} (va : ExSum sα a) (n : Q(ℕ)) : Result (ExSum sα) q($a ^ $n) := +partial def evalPowNat {a : Q($α)} (va : ExSum sα a) (n : Q(ℕ)) : + Lean.Core.CoreM <| Result (ExSum sα) q($a ^ $n) := do let nn := n.natLit! if nn = 1 then - ⟨_, va, (q(pow_one $a) : Expr)⟩ + return ⟨_, va, (q(pow_one $a) : Expr)⟩ else let nm := nn >>> 1 have m : Q(ℕ) := mkRawNatLit nm if nn &&& 1 = 0 then - let ⟨_, vb, pb⟩ := evalPowNat va m - let ⟨_, vc, pc⟩ := evalMul sα vb vb - ⟨_, vc, (q(pow_bit0 $pb $pc) : Expr)⟩ + let ⟨_, vb, pb⟩ ← evalPowNat va m + let ⟨_, vc, pc⟩ ← evalMul sα vb vb + return ⟨_, vc, (q(pow_bit0 $pb $pc) : Expr)⟩ else - let ⟨_, vb, pb⟩ := evalPowNat va m - let ⟨_, vc, pc⟩ := evalMul sα vb vb - let ⟨_, vd, pd⟩ := evalMul sα vc va - ⟨_, vd, (q(pow_bit1 $pb $pc $pd) : Expr)⟩ + let ⟨_, vb, pb⟩ ← evalPowNat va m + let ⟨_, vc, pc⟩ ← evalMul sα vb vb + let ⟨_, vd, pd⟩ ← evalMul sα vc va + return ⟨_, vd, (q(pow_bit1 $pb $pc $pd) : Expr)⟩ theorem one_pow (b : ℕ) : ((nat_lit 1).rawCast : R) ^ b = (nat_lit 1).rawCast := by simp @@ -750,10 +765,11 @@ theorem mul_pow {ea₁ b c₁ : ℕ} {xa₁ : R} In all other cases we use `evalPowProdAtom`. -/ def evalPowProd {a : Q($α)} {b : Q(ℕ)} (va : ExProd sα a) (vb : ExProd sℕ b) : - Result (ExProd sα) q($a ^ $b) := - let res : Option (Result (ExProd sα) q($a ^ $b)) := do + Lean.Core.CoreM <| Result (ExProd sα) q($a ^ $b) := do + Lean.Core.checkSystem decl_name%.toString + let res : OptionT Lean.Core.CoreM (Result (ExProd sα) q($a ^ $b)) := do match va, vb with - | .const 1, _ => some ⟨_, va, (q(one_pow (R := $α) $b) : Expr)⟩ + | .const 1, _ => return ⟨_, va, (q(one_pow (R := $α) $b) : Expr)⟩ | .const za ha, .const zb hb => assert! 0 ≤ zb let ra := Result.ofRawRat za a ha @@ -763,13 +779,13 @@ def evalPowProd {a : Q($α)} {b : Q(ℕ)} (va : ExProd sα a) (vb : ExProd sℕ q(CommSemiring.toSemiring) ra let ⟨zc, hc⟩ ← rc.toRatNZ let ⟨c, pc⟩ := rc.toRawEq - some ⟨c, .const zc hc, pc⟩ - | .mul vxa₁ vea₁ va₂, vb => do - let ⟨_, vc₁, pc₁⟩ := evalMulProd sℕ vea₁ vb - let ⟨_, vc₂, pc₂⟩ := evalPowProd va₂ vb - some ⟨_, .mul vxa₁ vc₁ vc₂, q(mul_pow $pc₁ $pc₂)⟩ - | _, _ => none - res.getD (evalPowProdAtom sα va vb) + return ⟨c, .const zc hc, pc⟩ + | .mul vxa₁ vea₁ va₂, vb => + let ⟨_, vc₁, pc₁⟩ ← evalMulProd sℕ vea₁ vb + let ⟨_, vc₂, pc₂⟩ ← evalPowProd va₂ vb + return ⟨_, .mul vxa₁ vc₁ vc₂, q(mul_pow $pc₁ $pc₂)⟩ + | _, _ => OptionT.fail + return (← res.run).getD (evalPowProdAtom sα va vb) /-- The result of `extractCoeff` is a numeral and a proof that the original expression @@ -827,24 +843,25 @@ theorem pow_nat {b c k : ℕ} {d e : R} (_ : b = c * k) (_ : a ^ c = d) (_ : d ^ Otherwise `a ^ b` is just encoded as `a ^ b * 1 + 0` using `evalPowAtom`. -/ partial def evalPow₁ {a : Q($α)} {b : Q(ℕ)} (va : ExSum sα a) (vb : ExProd sℕ b) : - Result (ExSum sα) q($a ^ $b) := + Lean.Core.CoreM <| Result (ExSum sα) q($a ^ $b) := do match va, vb with | va, .const 1 => haveI : $b =Q Nat.rawCast (nat_lit 1) := ⟨⟩ - ⟨_, va, q(pow_one_cast $a)⟩ + return ⟨_, va, q(pow_one_cast $a)⟩ | .zero, vb => match vb.evalPos with - | some p => ⟨_, .zero, q(zero_pow (R := $α) $p)⟩ - | none => evalPowAtom sα (.sum .zero) vb + | some p => return ⟨_, .zero, q(zero_pow (R := $α) $p)⟩ + | none => return evalPowAtom sα (.sum .zero) vb | ExSum.add va .zero, vb => -- TODO: using `.add` here takes a while to compile? - let ⟨_, vc, pc⟩ := evalPowProd sα va vb - ⟨_, vc.toSum, q(single_pow $pc)⟩ + let ⟨_, vc, pc⟩ ← evalPowProd sα va vb + return ⟨_, vc.toSum, q(single_pow $pc)⟩ | va, vb => if vb.coeff > 1 then let ⟨k, _, vc, pc⟩ := extractCoeff vb - let ⟨_, vd, pd⟩ := evalPow₁ va vc - let ⟨_, ve, pe⟩ := evalPowNat sα vd k - ⟨_, ve, q(pow_nat $pc $pd $pe)⟩ - else evalPowAtom sα (.sum va) vb + let ⟨_, vd, pd⟩ ← evalPow₁ va vc + let ⟨_, ve, pe⟩ ← evalPowNat sα vd k + return ⟨_, ve, q(pow_nat $pc $pd $pe)⟩ + else + return evalPowAtom sα (.sum va) vb theorem pow_zero (a : R) : a ^ 0 = (nat_lit 1).rawCast + 0 := by simp @@ -858,14 +875,14 @@ theorem pow_add {b₁ b₂ : ℕ} {d : R} * `a ^ (b₁ + b₂) = a ^ b₁ * a ^ b₂` -/ def evalPow {a : Q($α)} {b : Q(ℕ)} (va : ExSum sα a) (vb : ExSum sℕ b) : - Result (ExSum sα) q($a ^ $b) := + Lean.Core.CoreM <| Result (ExSum sα) q($a ^ $b) := do match vb with - | .zero => ⟨_, (ExProd.mkNat sα 1).2.toSum, q(pow_zero $a)⟩ + | .zero => return ⟨_, (ExProd.mkNat sα 1).2.toSum, q(pow_zero $a)⟩ | .add vb₁ vb₂ => - let ⟨_, vc₁, pc₁⟩ := evalPow₁ sα va vb₁ - let ⟨_, vc₂, pc₂⟩ := evalPow va vb₂ - let ⟨_, vd, pd⟩ := evalMul sα vc₁ vc₂ - ⟨_, vd, q(pow_add $pc₁ $pc₂ $pd)⟩ + let ⟨_, vc₁, pc₁⟩ ← evalPow₁ sα va vb₁ + let ⟨_, vc₂, pc₂⟩ ← evalPow va vb₂ + let ⟨_, vd, pd⟩ ← evalMul sα vc₁ vc₂ + return ⟨_, vd, q(pow_add $pc₁ $pc₂ $pd)⟩ /-- This cache contains data required by the `ring` tactic during execution. -/ structure Cache {α : Q(Type u)} (sα : Q(CommSemiring $α)) := @@ -968,6 +985,7 @@ def evalInvAtom (a : Q($α)) : AtomM (Result (ExBase sα) q($a⁻¹)) := do -/ def ExProd.evalInv {a : Q($α)} (czα : Option Q(CharZero $α)) (va : ExProd sα a) : AtomM (Result (ExProd sα) q($a⁻¹)) := do + Lean.Core.checkSystem decl_name%.toString match va with | .const c hc => let ra := Result.ofRawRat c a hc @@ -982,7 +1000,7 @@ def ExProd.evalInv {a : Q($α)} (czα : Option Q(CharZero $α)) (va : ExProd sα | .mul (x := a₁) (e := _a₂) _va₁ va₂ va₃ => do let ⟨_b₁, vb₁, pb₁⟩ ← evalInvAtom sα dα a₁ let ⟨_b₃, vb₃, pb₃⟩ ← va₃.evalInv czα - let ⟨c, vc, (pc : Q($_b₃ * ($_b₁ ^ $_a₂ * Nat.rawCast 1) = $c))⟩ := + let ⟨c, vc, (pc : Q($_b₃ * ($_b₁ ^ $_a₂ * Nat.rawCast 1) = $c))⟩ ← evalMulProd sα vb₃ (vb₁.toProd va₂) pure ⟨c, vc, (q(inv_mul $pb₁ $pb₃ $pc) : Expr)⟩ @@ -1014,7 +1032,7 @@ theorem div_pf {R} [DivisionRing R] {a b c d : R} (_ : b⁻¹ = c) (_ : a * c = def evalDiv {a b : Q($α)} (rα : Q(DivisionRing $α)) (czα : Option Q(CharZero $α)) (va : ExSum sα a) (vb : ExSum sα b) : AtomM (Result (ExSum sα) q($a / $b)) := do let ⟨_c, vc, pc⟩ ← vb.evalInv sα rα czα - let ⟨d, vd, (pd : Q($a * $_c = $d))⟩ := evalMul sα va vc + let ⟨d, vd, (pd : Q($a * $_c = $d))⟩ ← evalMul sα va vc pure ⟨d, vd, (q(div_pf $pc $pd) : Expr)⟩ theorem add_congr (_ : a = a') (_ : b = b') (_ : a' + b' = c) : (a + b : R) = c := by @@ -1089,14 +1107,14 @@ partial def eval {u : Lean.Level} {α : Q(Type u)} (sα : Q(CommSemiring $α)) | ~q($a + $b) => let ⟨_, va, pa⟩ ← eval sα c a let ⟨_, vb, pb⟩ ← eval sα c b - let ⟨c, vc, p⟩ := evalAdd sα va vb + let ⟨c, vc, p⟩ ← evalAdd sα va vb pure ⟨c, vc, (q(add_congr $pa $pb $p) : Expr)⟩ | _ => els | ``HMul.hMul, _, _ | ``Mul.mul, _, _ => match e with | ~q($a * $b) => let ⟨_, va, pa⟩ ← eval sα c a let ⟨_, vb, pb⟩ ← eval sα c b - let ⟨c, vc, p⟩ := evalMul sα va vb + let ⟨c, vc, p⟩ ← evalMul sα va vb pure ⟨c, vc, (q(mul_congr $pa $pb $p) : Expr)⟩ | _ => els | ``HSMul.hSMul, _, _ => match e with @@ -1110,19 +1128,20 @@ partial def eval {u : Lean.Level} {α : Q(Type u)} (sα : Q(CommSemiring $α)) | ~q($a ^ $b) => let ⟨_, va, pa⟩ ← eval sα c a let ⟨_, vb, pb⟩ ← eval sℕ .nat b - let ⟨c, vc, p⟩ := evalPow sα va vb + let ⟨c, vc, p⟩ ← evalPow sα va vb pure ⟨c, vc, (q(pow_congr $pa $pb $p) : Expr)⟩ | _ => els | ``Neg.neg, some rα, _ => match e with | ~q(-$a) => let ⟨_, va, pa⟩ ← eval sα c a - let ⟨b, vb, p⟩ := evalNeg sα rα va + let ⟨b, vb, p⟩ ← evalNeg sα rα va pure ⟨b, vb, (q(neg_congr $pa $p) : Expr)⟩ + | _ => els | ``HSub.hSub, some rα, _ | ``Sub.sub, some rα, _ => match e with | ~q($a - $b) => do let ⟨_, va, pa⟩ ← eval sα c a let ⟨_, vb, pb⟩ ← eval sα c b - let ⟨c, vc, p⟩ := evalSub sα rα va vb + let ⟨c, vc, p⟩ ← evalSub sα rα va vb pure ⟨c, vc, (q(sub_congr $pa $pb $p) : Expr)⟩ | _ => els | ``Inv.inv, _, some dα => match e with @@ -1130,6 +1149,7 @@ partial def eval {u : Lean.Level} {α : Q(Type u)} (sα : Q(CommSemiring $α)) let ⟨_, va, pa⟩ ← eval sα c a let ⟨b, vb, p⟩ ← va.evalInv sα dα c.czα pure ⟨b, vb, (q(inv_congr $pa $p) : Expr)⟩ + | _ => els | ``HDiv.hDiv, _, some dα | ``Div.div, _, some dα => match e with | ~q($a / $b) => do let ⟨_, va, pa⟩ ← eval sα c a diff --git a/test/tactic_timeout.lean b/test/tactic_timeout.lean new file mode 100644 index 0000000000000..455b2ff5e2399 --- /dev/null +++ b/test/tactic_timeout.lean @@ -0,0 +1,98 @@ +import Mathlib.Tactic.Linarith + +/-! +# Test that tactics respond to a cancellation request +-/ + + +variable {α} + +open Lean Elab Tactic + +/-! versions of try/catch that catch `interrupted` too -/ +section catch_interrupted +attribute [-instance] + Lean.instMonadExceptOfExceptionCoreM Lean.Elab.Tactic.instMonadExceptExceptionTacticM + +def Meta.tryCatchAll (m : MetaM α) (h : Exception → MetaM α) : MetaM α := tryCatch m h +def Term.tryCatchAll (m : TermElabM α) (h : Exception → TermElabM α) : TermElabM α := tryCatch m h +def Tactic.tryCatchAll (x : TacticM α) (h : Exception → TacticM α) : TacticM α := do + let b ← saveState + try x catch ex => b.restore; h ex + +end catch_interrupted + +section test_infra + +def Tactic.withTimeout (ms : UInt32) (t : TacticM α) : TacticM (α ⊕ Nat) := do + let tk ← IO.CancelToken.new + withTheReader Core.Context (fun s => { s with cancelTk? := some tk }) do + let t0 ← IO.monoMsNow + let watchdog ← IO.asTask do + IO.sleep ms + tk.set + let r ← Tactic.tryCatchAll (.inl <$> t) + (fun e => do + IO.cancel watchdog + if !e.isInterrupt || !(← tk.isSet) then + throw e + else + let duration := (← IO.monoMsNow) - t0 + return .inr duration) + IO.cancel watchdog + return r + +/-- `with_timeout 100 => tac` allows `tac` only 100ms to run. -/ +elab "with_timeout " ms:num "=>" tac:tacticSeq : tactic => do + let ms := ms.getNat.toUInt32 + if let .inr _duration ← Tactic.withTimeout ms (evalTactic tac) then + throwError f!"Tactic took more than {ms}ms" + +set_option linter.unusedTactic false + +/-- error: Tactic took more than 500ms -/ +#guard_msgs in +example : True := by + with_timeout 500 => + sleep 1000 + trivial + +example: True := by + with_timeout 500 => + sleep 100 + trivial + +end test_infra + +/-- `check_timeouts 100 => tac` checks that `tac` never goes longer than `100ms` without checking +for cancellation. -/ +elab "check_timeouts " tol_ms:num "=>" tac:tacticSeq : tactic => do + let mut t := 0 + let tol_ms := tol_ms.getNat + repeat do + if let .inr duration ← Tactic.withTimeout t.toUInt32 (evalTactic tac) then + if duration > t + tol_ms then + logError f!"Tactic took much more than {t}ms ({duration}ms)" + trace[debug] "Tactic overran from {t}ms to {duration}ms" + else + break + t := t + tol_ms + +set_option maxHeartbeats 0 +set_option linter.unusedTactic false +set_option linter.unusedVariables false + +theorem linear_combination_with_10_terms + (a b c d e f g h i j : Int) + (h0 : -e + g + -h + i = 0) + (h1 : b + -d + -e + f + g + i = 0) + (h2 : -b + j = 0) + (h3 : c + d + -f + -i = 0) + (h4 : b + c + e + -g + -h + i + j = 0) + (h5 : -a + b + d + f + -h + -i = 0) + (h6 : a + d + e + -g + -h = 0) + (h7 : -a + d + -f + -h + j = 0) + (h8 : a + -d + e + f + g + h + -i + j = 0) + (h9 : -a + b + c + -e + -f + h + j = 0) : + -2*a + b + 2*c + d + -3*f + -g + 3*h + -3*i = 0 := by + check_timeouts 250 => nlinarith From c8df482f4de3224efdb7bb67160f0af3ef3bc065 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 29 Sep 2024 04:48:27 +0000 Subject: [PATCH 033/174] chore: adaptation for batteries#905 (#17229) Co-authored-by: F. G. Dorais --- Mathlib/Data/List/Lemmas.lean | 4 ---- lake-manifest.json | 2 +- 2 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index 0d7c5ffd0b7ff..ea2d1faa23b35 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -38,10 +38,6 @@ theorem getLast_tail (l : List α) (hl : l.tail ≠ []) : · simp [show 1 + (l.length - 1 - 1) = l.length - 1 by omega] omega -lemma getElem_tail {i} (L : List α) (hi : i < L.tail.length) : - L.tail[i] = L[i + 1]'(by simp at *; omega) := by - induction L <;> simp at hi |- - @[deprecated (since := "2024-08-19")] alias nthLe_tail := getElem_tail theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : diff --git a/lake-manifest.json b/lake-manifest.json index 83f9c543ecef0..1aacaa9e2f075 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "51c38e3828d06c82741a7a65df93611e2ce209e1", + "rev": "bf12ff6041cbab6eba6b54d9467baed807bb2bfd", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From ae458abd0d248cbddff203c64074f5bb15331c80 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 29 Sep 2024 06:55:53 +0000 Subject: [PATCH 034/174] chore: robustify usage of List.prod API (#17220) Remove some reliances on defeq, and use some named arguments. I'm hoping to subsequently change the defeq, in https://github.com/leanprover-community/mathlib4/pull/17223 --- Mathlib/Algebra/BigOperators/Group/List.lean | 25 +++++++------------- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 2 +- Mathlib/Analysis/Normed/Field/Basic.lean | 2 +- Mathlib/Data/List/Prime.lean | 4 ++-- Mathlib/Deprecated/Subgroup.lean | 3 ++- Mathlib/Deprecated/Subring.lean | 3 ++- Mathlib/GroupTheory/FixedPointFree.lean | 4 ++-- Mathlib/GroupTheory/FreeGroup/Basic.lean | 4 ++-- 8 files changed, 21 insertions(+), 26 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 00a2158bc6981..f097934c26635 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -189,21 +189,14 @@ theorem prod_isUnit_iff {α : Type*} [CommMonoid α] {L : List α} : exact fun m' h' ↦ Or.elim (eq_or_mem_of_mem_cons h') (fun H => H.substr h.1) fun H => ih h.2 _ H @[to_additive (attr := simp)] -theorem prod_take_mul_prod_drop : ∀ (L : List M) (i : ℕ), (L.take i).prod * (L.drop i).prod = L.prod - | [], i => by simp [Nat.zero_le] - | L, 0 => by simp - | h :: t, n + 1 => by - dsimp - rw [prod_cons, prod_cons, mul_assoc, prod_take_mul_prod_drop t] +theorem prod_take_mul_prod_drop (L : List M) (i : ℕ) : + (L.take i).prod * (L.drop i).prod = L.prod := by + simp [← prod_append] @[to_additive (attr := simp)] -theorem prod_take_succ : - ∀ (L : List M) (i : ℕ) (p : i < L.length), (L.take (i + 1)).prod = (L.take i).prod * L[i] - | [], i, p => by cases p - | h :: t, 0, _ => rfl - | h :: t, n + 1, p => by - dsimp - rw [prod_cons, prod_cons, prod_take_succ t n (Nat.lt_of_succ_lt_succ p), mul_assoc] +theorem prod_take_succ (L : List M) (i : ℕ) (p : i < L.length) : + (L.take (i + 1)).prod = (L.take i).prod * L[i] := by + simp [take_succ, p] /-- A list with product not one must have positive length. -/ @[to_additive "A list with sum not zero must have positive length."] @@ -273,9 +266,9 @@ last. -/ @[to_additive "A variant of `sum_range_succ` which pulls off the first term in the sum rather than the last."] lemma prod_range_succ' (f : ℕ → M) (n : ℕ) : - ((range n.succ).map f).prod = f 0 * ((range n).map fun i ↦ f i.succ).prod := - Nat.recOn n (show 1 * f 0 = f 0 * 1 by rw [one_mul, mul_one]) fun _ hd => by - rw [List.prod_range_succ, hd, mul_assoc, ← List.prod_range_succ] + ((range n.succ).map f).prod = f 0 * ((range n).map fun i ↦ f i.succ).prod := by + rw [range_succ_eq_map] + simp [Function.comp_def] @[to_additive] lemma prod_eq_one (hl : ∀ x ∈ l, x = 1) : l.prod = 1 := by induction l with diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 03dc4ba71c416..2c20870f69901 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -791,7 +791,7 @@ theorem mem_closure_iff_exists_list {R} [Semiring R] {s : Set R} {x} : ⟨[t], List.forall_mem_singleton.2 ht1, by rw [List.map_singleton, List.sum_singleton, ht2]⟩ Submonoid.closure_induction hx - (fun x hx => ⟨[x], List.forall_mem_singleton.2 hx, one_mul x⟩) + (fun x hx => ⟨[x], List.forall_mem_singleton.2 hx, List.prod_singleton⟩) ⟨[], List.forall_mem_nil _, rfl⟩ fun x y ⟨t, ht1, ht2⟩ ⟨u, hu1, hu2⟩ => ⟨t ++ u, List.forall_mem_append.2 ⟨ht1, hu1⟩, by rw [List.prod_append, ht2, hu2]⟩) ⟨[], List.forall_mem_nil _, rfl⟩ fun x y ⟨L, HL1, HL2⟩ ⟨M, HM1, HM2⟩ => diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 56b1e27ee82ae..a8762297c678a 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -348,7 +348,7 @@ theorem List.norm_prod_le' : ∀ {l : List α}, l ≠ [] → ‖l.prod‖ ≤ (l | [], h => (h rfl).elim | [a], _ => by simp | a::b::l, _ => by - rw [List.map_cons, List.prod_cons, @List.prod_cons _ _ _ ‖a‖] + rw [List.map_cons, List.prod_cons, List.prod_cons (a := ‖a‖)] refine le_trans (norm_mul_le _ _) (mul_le_mul_of_nonneg_left ?_ (norm_nonneg _)) exact List.norm_prod_le' (List.cons_ne_nil b l) diff --git a/Mathlib/Data/List/Prime.lean b/Mathlib/Data/List/Prime.lean index e5812cb86b262..67c3197c12c0c 100644 --- a/Mathlib/Data/List/Prime.lean +++ b/Mathlib/Data/List/Prime.lean @@ -54,10 +54,10 @@ theorem perm_of_prod_eq_prod : ∀ {l₁ l₂ : List M}, l₁.prod = l₂.prod → (∀ p ∈ l₁, Prime p) → (∀ p ∈ l₂, Prime p) → Perm l₁ l₂ | [], [], _, _, _ => Perm.nil | [], a :: l, h₁, _, h₃ => - have ha : a ∣ 1 := @prod_nil M _ ▸ h₁.symm ▸ (@prod_cons _ _ l a).symm ▸ dvd_mul_right _ _ + have ha : a ∣ 1 := prod_nil (M := M) ▸ h₁.symm ▸ (prod_cons (l := l)).symm ▸ dvd_mul_right _ _ absurd ha (Prime.not_dvd_one (h₃ a (mem_cons_self _ _))) | a :: l, [], h₁, h₂, _ => - have ha : a ∣ 1 := @prod_nil M _ ▸ h₁ ▸ (@prod_cons _ _ l a).symm ▸ dvd_mul_right _ _ + have ha : a ∣ 1 := prod_nil (M := M) ▸ h₁ ▸ (prod_cons (l := l)).symm ▸ dvd_mul_right _ _ absurd ha (Prime.not_dvd_one (h₂ a (mem_cons_self _ _))) | a :: l₁, b :: l₂, h, hl₁, hl₂ => by classical diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index b7c3f4795f8ec..f30ed0f0e3ca3 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -453,7 +453,8 @@ theorem closure_subgroup {s : Set G} (hs : IsSubgroup s) : closure s = s := @[to_additive] theorem exists_list_of_mem_closure {s : Set G} {a : G} (h : a ∈ closure s) : ∃ l : List G, (∀ x ∈ l, x ∈ s ∨ x⁻¹ ∈ s) ∧ l.prod = a := - InClosure.recOn h (fun {x} hxs => ⟨[x], List.forall_mem_singleton.2 <| Or.inl hxs, one_mul _⟩) + InClosure.recOn h + (fun {x} hxs => ⟨[x], List.forall_mem_singleton.2 <| Or.inl hxs, List.prod_singleton⟩) ⟨[], List.forall_mem_nil _, rfl⟩ (fun {x} _ ⟨L, HL1, HL2⟩ => ⟨L.reverse.map Inv.inv, fun x hx => diff --git a/Mathlib/Deprecated/Subring.lean b/Mathlib/Deprecated/Subring.lean index 7c267c4d4f513..20fc691ea4b93 100644 --- a/Mathlib/Deprecated/Subring.lean +++ b/Mathlib/Deprecated/Subring.lean @@ -97,7 +97,8 @@ theorem exists_list_of_mem_closure {a : R} (h : a ∈ closure s) : ∃ L : List (List R), (∀ l ∈ L, ∀ x ∈ l, x ∈ s ∨ x = (-1 : R)) ∧ (L.map List.prod).sum = a := AddGroup.InClosure.recOn h fun {x} hx ↦ match x, Monoid.exists_list_of_mem_closure hx with - | _, ⟨L, h1, rfl⟩ => ⟨[L], List.forall_mem_singleton.2 fun r hr ↦ Or.inl (h1 r hr), zero_add _⟩ + | _, ⟨L, h1, rfl⟩ => + ⟨[L], List.forall_mem_singleton.2 fun r hr ↦ Or.inl (h1 r hr), List.sum_singleton⟩ ⟨[], List.forall_mem_nil _, rfl⟩ fun {b} _ ih ↦ match b, ih with | _, ⟨L1, h1, rfl⟩ => diff --git a/Mathlib/GroupTheory/FixedPointFree.lean b/Mathlib/GroupTheory/FixedPointFree.lean index 2f9fb7c2c40f9..a91326561faf5 100644 --- a/Mathlib/GroupTheory/FixedPointFree.lean +++ b/Mathlib/GroupTheory/FixedPointFree.lean @@ -54,8 +54,8 @@ theorem prod_pow_eq_one (hφ : FixedPointFree φ) {n : ℕ} (hn : φ^[n] = _root theorem coe_eq_inv_of_sq_eq_one (hφ : FixedPointFree φ) (h2 : φ^[2] = _root_.id) : ⇑φ = (·⁻¹) := by ext g - have key : 1 * g * φ g = 1 := hφ.prod_pow_eq_one h2 g - rwa [one_mul, ← inv_eq_iff_mul_eq_one, eq_comm] at key + have key : g * φ g = 1 := by simpa [List.range_succ] using hφ.prod_pow_eq_one h2 g + rwa [← inv_eq_iff_mul_eq_one, eq_comm] at key section Involutive diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 0718b7ca18057..e1096e6829211 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -571,7 +571,7 @@ def lift : (α → β) ≃ (FreeGroup α →* β) where MonoidHom.mk' (Quot.lift (Lift.aux f) fun L₁ L₂ => Red.Step.lift) <| by rintro ⟨L₁⟩ ⟨L₂⟩; simp [Lift.aux] invFun g := g ∘ of - left_inv f := one_mul _ + left_inv f := List.prod_singleton right_inv g := MonoidHom.ext <| by rintro ⟨L⟩ @@ -592,7 +592,7 @@ theorem lift.mk : lift f (mk L) = List.prod (L.map fun x => cond x.2 (f x.1) (f @[to_additive (attr := simp)] theorem lift.of {x} : lift f (of x) = f x := - one_mul _ + List.prod_singleton @[to_additive] theorem lift.unique (g : FreeGroup α →* β) (hg : ∀ x, g (FreeGroup.of x) = f x) {x} : From 9491960f1f448fa353e840f5b146feea22d5dbfd Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sun, 29 Sep 2024 08:12:12 +0000 Subject: [PATCH 035/174] feat(Topology/Group/Profinite): limits in the category of profinite groups (#16976) Define the limit in profinite groups as a profinite group. Co-authored-by: Jujian Zhang Nailin Guan <3571819596@qq.com> Co-authored-by: dagurtomas --- Mathlib.lean | 2 +- .../Basic.lean} | 96 ++++++++++++++++++- 2 files changed, 93 insertions(+), 5 deletions(-) rename Mathlib/Topology/Algebra/Category/{ProfiniteGrp.lean => ProfiniteGrp/Basic.lean} (60%) diff --git a/Mathlib.lean b/Mathlib.lean index 1431bd8d7f883..5e6c98ddc33a8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4408,7 +4408,7 @@ import Mathlib.Topology.AlexandrovDiscrete import Mathlib.Topology.Algebra.Affine import Mathlib.Topology.Algebra.Algebra import Mathlib.Topology.Algebra.Algebra.Rat -import Mathlib.Topology.Algebra.Category.ProfiniteGrp +import Mathlib.Topology.Algebra.Category.ProfiniteGrp.Basic import Mathlib.Topology.Algebra.ConstMulAction import Mathlib.Topology.Algebra.Constructions import Mathlib.Topology.Algebra.Constructions.DomMulAct diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean similarity index 60% rename from Mathlib/Topology/Algebra/Category/ProfiniteGrp.lean rename to Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index 7f613f00bfb58..58f632d59317e 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -6,8 +6,6 @@ Authors: Jujian Zhang, Nailin Guan, Yuyang Zhao import Mathlib.Algebra.Category.Grp.FiniteGrp import Mathlib.Topology.Algebra.ContinuousMonoidHom import Mathlib.Topology.Category.Profinite.Basic - - /-! # Category of Profinite Groups @@ -23,10 +21,9 @@ disconnected. * `ofFiniteGrp` : A `FiniteGrp` when given the discrete topology can be considered as a profinite group. - -/ -universe u +universe u v open CategoryTheory Topology @@ -161,4 +158,95 @@ instance : HasForget₂ ProfiniteGrp Grp where map := fun f => f.toMonoidHom } +/-- The functor mapping a profinite group to its underlying profinite space. -/ +def profiniteGrpToProfinite : ProfiniteGrp ⥤ Profinite where + obj G := G.toProfinite + map f := ⟨f, by continuity⟩ + +instance : profiniteGrpToProfinite.Faithful := { + map_injective := fun {_ _} _ _ h => + ConcreteCategory.hom_ext_iff.mpr (congrFun (congrArg ContinuousMap.toFun h)) } + end ProfiniteGrp + +/-! +# Limits in the category of profinite groups + +In this section, we construct limits in the category of profinite groups. + +* `ProfiniteGrp.limitCone` : The explicit limit cone in `ProfiniteGrp`. + +* `ProfiniteGrp.limitConeIsLimit`: `ProfiniteGrp.limitCone` is a limit cone. + +-/ + +section Limits + +namespace ProfiniteGrp + +section + +variable {J : Type v} [SmallCategory J] (F : J ⥤ ProfiniteGrp.{max v u}) + +/-- Auxiliary construction to obtain the group structure on the limit of profinite groups. -/ +def limitConePtAux : Subgroup (Π j : J, F.obj j) where + carrier := {x | ∀ ⦃i j : J⦄ (π : i ⟶ j), F.map π (x i) = x j} + mul_mem' hx hy _ _ π := by simp only [Pi.mul_apply, map_mul, hx π, hy π] + one_mem' := by simp only [Set.mem_setOf_eq, Pi.one_apply, map_one, implies_true] + inv_mem' h _ _ π := by simp only [Pi.inv_apply, map_inv, h π] + +instance : Group (Profinite.limitCone (F ⋙ profiniteGrpToProfinite.{max v u})).pt := + inferInstanceAs (Group (limitConePtAux F)) + +instance : TopologicalGroup (Profinite.limitCone (F ⋙ profiniteGrpToProfinite.{max v u})).pt := + inferInstanceAs (TopologicalGroup (limitConePtAux F)) + +/-- The explicit limit cone in `ProfiniteGrp`. -/ +abbrev limitCone : Limits.Cone F where + pt := ofProfinite (Profinite.limitCone (F ⋙ profiniteGrpToProfinite.{max v u})).pt + π := + { app := fun j => { + toFun := fun x => x.1 j + map_one' := rfl + map_mul' := fun x y => rfl + continuous_toFun := by + exact (continuous_apply j).comp (continuous_iff_le_induced.mpr fun U a => a) } + naturality := fun i j f => by + simp only [Functor.const_obj_obj, Functor.comp_obj, + Functor.const_obj_map, Category.id_comp, Functor.comp_map] + congr + exact funext fun x => (x.2 f).symm } + +/-- `ProfiniteGrp.limitCone` is a limit cone. -/ +def limitConeIsLimit : Limits.IsLimit (limitCone F) where + lift cone := { + ((Profinite.limitConeIsLimit (F ⋙ profiniteGrpToProfinite)).lift + (profiniteGrpToProfinite.mapCone cone)) with + map_one' := Subtype.ext (funext fun j ↦ map_one (cone.π.app j)) + -- TODO: investigate whether it's possible to set up `ext` lemmas for the `TopCat`-related + -- categories so that `by ext j; exact map_one (cone.π.app j)` works here, similarly below. + map_mul' := fun _ _ ↦ Subtype.ext (funext fun j ↦ map_mul (cone.π.app j) _ _) } + uniq cone m h := by + apply profiniteGrpToProfinite.map_injective + simpa using (Profinite.limitConeIsLimit (F ⋙ profiniteGrpToProfinite)).uniq + (profiniteGrpToProfinite.mapCone cone) (profiniteGrpToProfinite.map m) + (fun j ↦ congrArg profiniteGrpToProfinite.map (h j)) + +instance : Limits.HasLimit F where + exists_limit := Nonempty.intro + { cone := limitCone F + isLimit := limitConeIsLimit F } + +/-- The abbreviation for the limit of `ProfiniteGrp`s. -/ +abbrev limit : ProfiniteGrp := (ProfiniteGrp.limitCone F).pt + +end + +instance : Limits.PreservesLimits profiniteGrpToProfinite.{u} where + preservesLimitsOfShape := { + preservesLimit := fun {F} ↦ CategoryTheory.Limits.preservesLimitOfPreservesLimitCone + (limitConeIsLimit F) (Profinite.limitConeIsLimit (F ⋙ profiniteGrpToProfinite)) } + +end ProfiniteGrp + +end Limits From a6006d2f1d3d5ffc24da15256c54ff36b26d9f61 Mon Sep 17 00:00:00 2001 From: Marcus Rossel Date: Sun, 29 Sep 2024 08:52:34 +0000 Subject: [PATCH 036/174] feat: stability of `List.insertionSort` (#16065) The statements of stability follow those used for `mergeSort` in leanprover/lean4#5092. Co-authored-by: Matthew Ballard --- Mathlib/Data/List/Sort.lean | 87 +++++++++++++++++++++++++++++++++++++ 1 file changed, 87 insertions(+) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 6936029a0217b..9742e4370e93c 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -345,6 +345,40 @@ theorem sublist_orderedInsert (x : α) (xs : List α) : xs <+ xs.orderedInsert r refine Sublist.trans ?_ (.append_left (.cons _ (.refl _)) _) rw [takeWhile_append_dropWhile] +theorem cons_sublist_orderedInsert {l c : List α} {a : α} (hl : c <+ l) (ha : ∀ a' ∈ c, a ≼ a') : + a :: c <+ orderedInsert r a l := by + induction l with + | nil => simp_all only [sublist_nil, orderedInsert, Sublist.refl] + | cons _ _ ih => + unfold orderedInsert + split_ifs with hr + · exact .cons₂ _ hl + · cases hl with + | cons _ h => exact .cons _ <| ih h + | cons₂ => exact absurd (ha _ <| mem_cons_self ..) hr + +theorem Sublist.orderedInsert_sublist [IsTrans α r] {as bs} (x) (hs : as <+ bs) (hb : bs.Sorted r) : + orderedInsert r x as <+ orderedInsert r x bs := by + cases as with + | nil => simp + | cons a as => + cases bs with + | nil => contradiction + | cons b bs => + unfold orderedInsert + cases hs <;> split_ifs with hr + · exact .cons₂ _ <| .cons _ ‹a :: as <+ bs› + · have ih := orderedInsert_sublist x ‹a :: as <+ bs› hb.of_cons + simp only [hr, orderedInsert, ite_true] at ih + exact .trans ih <| .cons _ (.refl _) + · have hba := pairwise_cons.mp hb |>.left _ (mem_of_cons_sublist ‹a :: as <+ bs›) + exact absurd (trans_of _ ‹r x b› hba) hr + · have ih := orderedInsert_sublist x ‹a :: as <+ bs› hb.of_cons + rw [orderedInsert, if_neg hr] at ih + exact .cons _ ih + · simp_all only [sorted_cons, cons_sublist_cons] + · exact .cons₂ _ <| orderedInsert_sublist x ‹as <+ bs› hb.of_cons + section TotalAndTransitive variable [IsTotal α r] [IsTrans α r] @@ -374,6 +408,59 @@ theorem sorted_insertionSort : ∀ l, Sorted r (insertionSort r l) end TotalAndTransitive +/-- +If `c` is a sorted sublist of `l`, then `c` is still a sublist of `insertionSort r l`. +-/ +theorem sublist_insertionSort {l c : List α} (hr : c.Pairwise r) (hc : c <+ l) : + c <+ insertionSort r l := by + induction l generalizing c with + | nil => simp_all only [sublist_nil, insertionSort, Sublist.refl] + | cons _ _ ih => + cases hc with + | cons _ h => exact ih hr h |>.trans (sublist_orderedInsert ..) + | cons₂ _ h => + obtain ⟨hr, hp⟩ := pairwise_cons.mp hr + exact cons_sublist_orderedInsert (ih hp h) hr + +/-- +Another statement of stability of insertion sort. +If a pair `[a, b]` is a sublist of `l` and `r a b`, +then `[a, b]` is still a sublist of `insertionSort r l`. +-/ +theorem pair_sublist_insertionSort {a b : α} {l : List α} (hab : r a b) (h : [a, b] <+ l) : + [a, b] <+ insertionSort r l := + sublist_insertionSort (pairwise_pair.mpr hab) h + +variable [IsAntisymm α r] [IsTotal α r] [IsTrans α r] + +/-- +A version of `insertionSort_stable` which only assumes `c <+~ l` (instead of `c <+ l`), but +additionally requires `IsAntisymm α r`, `IsTotal α r` and `IsTrans α r`. +-/ +theorem sublist_insertionSort' {l c : List α} (hs : c.Sorted r) (hc : c <+~ l) : + c <+ insertionSort r l := by + classical + obtain ⟨d, hc, hd⟩ := hc + induction l generalizing c d with + | nil => simp_all only [sublist_nil, insertionSort, nil_perm] + | cons a _ ih => + cases hd with + | cons _ h => exact ih hs _ hc h |>.trans (sublist_orderedInsert ..) + | cons₂ _ h => + specialize ih (hs.erase _) _ (erase_cons_head a ‹List _› ▸ hc.erase a) h + have hm := hc.mem_iff.mp <| mem_cons_self .. + have he := orderedInsert_erase _ _ hm hs + exact he ▸ Sublist.orderedInsert_sublist _ ih (sorted_insertionSort ..) + +/-- +Another statement of stability of insertion sort. +If a pair `[a, b]` is a sublist of a permutation of `l` and `a ≼ b`, +then `[a, b]` is still a sublist of `insertionSort r l`. +-/ +theorem pair_sublist_insertionSort' {a b : α} {l : List α} (hab : a ≼ b) (h : [a, b] <+~ l) : + [a, b] <+ insertionSort r l := + sublist_insertionSort' (pairwise_pair.mpr hab) h + end Correctness end InsertionSort From 62eccbcb33b8a9fc8e7cc61cd546298b52f0c29b Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sun, 29 Sep 2024 12:03:25 +0000 Subject: [PATCH 037/174] chore: update Mathlib dependencies 2024-09-29 (#17252) 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 1aacaa9e2f075..3e190c4de8899 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a895713f7701e295a015b1087f3113fd3d615272", + "rev": "50aaaf78b7db5bd635c19c660d59ed31b9bc9b5a", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From e17fd83dc9062e01ffc4e234c11db0bbe7cc277f Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sun, 29 Sep 2024 13:27:48 +0000 Subject: [PATCH 038/174] feat(Topology/Group): define closed subgroup (#16977) Define the type of closed subgroup in a topological group. --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 7 + .../Algebra/Category/ProfiniteGrp/Basic.lean | 9 ++ Mathlib/Topology/Algebra/ClosedSubgroup.lean | 128 ++++++++++++++++++ 4 files changed, 145 insertions(+) create mode 100644 Mathlib/Topology/Algebra/ClosedSubgroup.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5e6c98ddc33a8..8ddcc87510e87 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4409,6 +4409,7 @@ import Mathlib.Topology.Algebra.Affine import Mathlib.Topology.Algebra.Algebra import Mathlib.Topology.Algebra.Algebra.Rat import Mathlib.Topology.Algebra.Category.ProfiniteGrp.Basic +import Mathlib.Topology.Algebra.ClosedSubgroup import Mathlib.Topology.Algebra.ConstMulAction import Mathlib.Topology.Algebra.Constructions import Mathlib.Topology.Algebra.Constructions.DomMulAct diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 17dc2929357ae..99bb82ddad9cc 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -392,6 +392,13 @@ theorem Normal.conjAct {G : Type*} [Group G] {H : Subgroup G} (hH : H.Normal) (g theorem smul_normal (g : G) (H : Subgroup G) [h : Normal H] : MulAut.conj g • H = H := h.conjAct g +theorem normalCore_eq_iInf_conjAct (H : Subgroup G) : + H.normalCore = ⨅ (g : ConjAct G), g • H := by + ext g + simp only [Subgroup.normalCore, Subgroup.mem_iInf, Subgroup.mem_pointwise_smul_iff_inv_smul_mem] + refine ⟨fun h x ↦ h x⁻¹, fun h x ↦ ?_⟩ + simpa only [ConjAct.toConjAct_inv, inv_inv] using h x⁻¹ + end Group section GroupWithZero diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index 58f632d59317e..803016fe59879 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Nailin Guan, Yuyang Zhao -/ import Mathlib.Algebra.Category.Grp.FiniteGrp +import Mathlib.Topology.Algebra.ClosedSubgroup import Mathlib.Topology.Algebra.ContinuousMonoidHom import Mathlib.Topology.Category.Profinite.Basic /-! @@ -21,6 +22,9 @@ disconnected. * `ofFiniteGrp` : A `FiniteGrp` when given the discrete topology can be considered as a profinite group. + +* `ofClosedSubgroup` : A closed subgroup of a profinite group is profinite. + -/ universe u v @@ -158,6 +162,11 @@ instance : HasForget₂ ProfiniteGrp Grp where map := fun f => f.toMonoidHom } +/-- A closed subgroup of a profinite group is profinite. -/ +def ofClosedSubgroup {G : ProfiniteGrp} (H : ClosedSubgroup G) : ProfiniteGrp := + letI : CompactSpace H := inferInstance + of H.1 + /-- The functor mapping a profinite group to its underlying profinite space. -/ def profiniteGrpToProfinite : ProfiniteGrp ⥤ Profinite where obj G := G.toProfinite diff --git a/Mathlib/Topology/Algebra/ClosedSubgroup.lean b/Mathlib/Topology/Algebra/ClosedSubgroup.lean new file mode 100644 index 0000000000000..302f09738270a --- /dev/null +++ b/Mathlib/Topology/Algebra/ClosedSubgroup.lean @@ -0,0 +1,128 @@ +/- +Copyright (c) 2024 Nailin Guan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nailin Guan +-/ + +import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.GroupTheory.Index + +/-! +# Closed subgroups of a topological group + +This files builds the SemilatticeInf `ClosedSubgroup G` of closed subgroups in a +topological group `G`, and its additive version `ClosedAddSubgroup`. + +# Main definitions and results + +* `normalCore_isClosed` : The `normalCore` of a closed subgroup is closed. + +* `finindex_closedSubgroup_isOpen` : A closed subgroup with finite index is open. + +-/ + +section + +universe u v + +/-- The type of closed subgroups of a topological group. -/ +@[ext] +structure ClosedSubgroup (G : Type u) [Group G] [TopologicalSpace G] extends Subgroup G where + isClosed' : IsClosed carrier + +/-- The type of closed subgroups of an additive topological group. -/ +@[ext] +structure ClosedAddSubgroup (G : Type u) [AddGroup G] [TopologicalSpace G] extends + AddSubgroup G where + isClosed' : IsClosed carrier + +attribute [to_additive] ClosedSubgroup + +attribute [coe] ClosedSubgroup.toSubgroup ClosedAddSubgroup.toAddSubgroup + +namespace ClosedSubgroup + +variable (G : Type u) [Group G] [TopologicalSpace G] + +variable {G} in +@[to_additive] +theorem toSubgroup_injective : Function.Injective + (ClosedSubgroup.toSubgroup : ClosedSubgroup G → Subgroup G) := + fun A B h ↦ by + ext + rw [h] + +@[to_additive] +instance : SetLike (ClosedSubgroup G) G where + coe U := U.1 + coe_injective' _ _ h := toSubgroup_injective <| SetLike.ext' h + +@[to_additive] +instance : SubgroupClass (ClosedSubgroup G) G where + mul_mem := Subsemigroup.mul_mem' _ + one_mem U := U.one_mem' + inv_mem := Subgroup.inv_mem' _ + +@[to_additive] +instance : Coe (ClosedSubgroup G) (Subgroup G) where + coe := toSubgroup + +@[to_additive] +instance instInfClosedSubgroup : Inf (ClosedSubgroup G) := + ⟨fun U V ↦ ⟨U ⊓ V, U.isClosed'.inter V.isClosed'⟩⟩ + +@[to_additive] +instance instSemilatticeInfClosedSubgroup : SemilatticeInf (ClosedSubgroup G) := + SetLike.coe_injective.semilatticeInf ((↑) : ClosedSubgroup G → Set G) fun _ _ ↦ rfl + +@[to_additive] +instance [CompactSpace G] (H : ClosedSubgroup G) : CompactSpace H := + isCompact_iff_compactSpace.mp (IsClosed.isCompact H.isClosed') + +end ClosedSubgroup + +open scoped Pointwise + +namespace Subgroup + +variable {G : Type u} [Group G] [TopologicalSpace G] [ContinuousMul G] + +lemma normalCore_isClosed (H : Subgroup G) (h : IsClosed (H : Set G)) : + IsClosed (H.normalCore : Set G) := by + rw [normalCore_eq_iInf_conjAct] + push_cast + apply isClosed_iInter + intro g + convert IsClosed.preimage (TopologicalGroup.continuous_conj (ConjAct.ofConjAct g⁻¹)) h + exact Set.ext (fun t ↦ Set.mem_smul_set_iff_inv_smul_mem) + +@[to_additive] +lemma isOpen_of_isClosed_of_finiteIndex (H : Subgroup G) [H.FiniteIndex] + (h : IsClosed (H : Set G)) : IsOpen (H : Set G) := by + apply isClosed_compl_iff.mp + convert isClosed_iUnion_of_finite <| fun (x : {x : (G ⧸ H) // x ≠ QuotientGroup.mk 1}) + ↦ IsClosed.smul h (Quotient.out' x.1) + ext x + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · have : QuotientGroup.mk 1 ≠ QuotientGroup.mk (s := H) x := by + apply QuotientGroup.eq.not.mpr + simpa only [inv_one, one_mul, ne_eq] + simp only [ne_eq, Set.mem_iUnion] + use ⟨QuotientGroup.mk (s := H) x, this.symm⟩, + (Quotient.out' (QuotientGroup.mk (s := H) x))⁻¹ * x + simp only [SetLike.mem_coe, smul_eq_mul, mul_inv_cancel_left, and_true] + exact QuotientGroup.eq.mp <| QuotientGroup.out_eq' (QuotientGroup.mk (s := H) x) + · rcases h with ⟨S,⟨y,hS⟩,mem⟩ + simp only [← hS] at mem + rcases mem with ⟨h,hh,eq⟩ + simp only [Set.mem_compl_iff, SetLike.mem_coe] + by_contra mH + simp only [← eq, ne_eq, smul_eq_mul] at mH + absurd y.2.symm + rw [← QuotientGroup.out_eq' y.1, QuotientGroup.eq] + simp only [inv_one, ne_eq, one_mul, (Subgroup.mul_mem_cancel_right H hh).mp mH] + +end Subgroup + +end From cda107d63d2ef806382bb7dfd92ac681500a1b3c Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sun, 29 Sep 2024 14:29:38 +0000 Subject: [PATCH 039/174] feat(RingTheory/Localization): some API on localizations (#17030) This PR adds some API for localizations of rings. In particular: - localizations commute with localizations - localizing away from a product of elements is the same as successive localization at the elements --- .../Morphisms/QuasiCompact.lean | 2 +- .../RingTheory/Localization/Away/Basic.lean | 102 +++++++++++++++++- Mathlib/RingTheory/Localization/Basic.lean | 42 ++++++++ 3 files changed, 144 insertions(+), 2 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 211a07cc5edf0..c85901ed45bc9 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -198,7 +198,7 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch {U : X.Opens} (hU : IsAffineOpen U) (x f : Γ(X, U)) (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by rw [← map_zero (X.presheaf.map (homOfLE <| X.basicOpen_le f : X.basicOpen f ⟶ U).op)] at H - obtain ⟨⟨_, n, rfl⟩, e⟩ := (hU.isLocalization_basicOpen f).exists_of_eq H + obtain ⟨n, e⟩ := (hU.isLocalization_basicOpen f).exists_of_eq H exact ⟨n, by simpa [mul_comm x] using e⟩ /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index c41d650a75ef4..eaaebda203892 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -38,7 +38,8 @@ section Away variable (x : R) /-- Given `x : R`, the typeclass `IsLocalization.Away x S` states that `S` is -isomorphic to the localization of `R` at the submonoid generated by `x`. -/ +isomorphic to the localization of `R` at the submonoid generated by `x`. +See `IsLocalization.Away.mk` for a specialized constructor. -/ abbrev Away (S : Type*) [CommSemiring S] [Algebra R S] := IsLocalization (Submonoid.powers x) S @@ -68,6 +69,58 @@ lemma sec_spec (s : S) : s * (algebraMap R S) (x ^ (IsLocalization.Away.sec x s) congr exact (IsLocalization.sec (Submonoid.powers x) s).2.property.choose_spec +lemma algebraMap_pow_isUnit (n : ℕ) : IsUnit (algebraMap R S x ^ n) := + IsUnit.pow _ <| IsLocalization.map_units _ (⟨x, 1, by simp⟩ : Submonoid.powers x) + +lemma algebraMap_isUnit : IsUnit (algebraMap R S x) := + IsLocalization.map_units _ (⟨x, 1, by simp⟩ : Submonoid.powers x) + +lemma surj (z : S) : ∃ (n : ℕ) (a : R), z * algebraMap R S x ^ n = algebraMap R S a := by + obtain ⟨⟨a, ⟨-, n, rfl⟩⟩, h⟩ := IsLocalization.surj (Submonoid.powers x) z + use n, a + simpa using h + +lemma exists_of_eq {a b : R} (h : algebraMap R S a = algebraMap R S b) : + ∃ (n : ℕ), x ^ n * a = x ^ n * b := by + obtain ⟨⟨-, n, rfl⟩, hx⟩ := IsLocalization.exists_of_eq (M := Submonoid.powers x) h + use n + +/-- Specialized constructor for `IsLocalization.Away`. -/ +lemma mk (r : R) (map_unit : IsUnit (algebraMap R S r)) + (surj : ∀ s, ∃ (n : ℕ) (a : R), s * algebraMap R S r ^ n = algebraMap R S a) + (exists_of_eq : ∀ a b, algebraMap R S a = algebraMap R S b → ∃ (n : ℕ), r ^ n * a = r ^ n * b) : + IsLocalization.Away r S where + map_units' := by + rintro ⟨-, n, rfl⟩ + simp only [map_pow] + exact IsUnit.pow _ map_unit + surj' z := by + obtain ⟨n, a, hn⟩ := surj z + use ⟨a, ⟨r ^ n, n, rfl⟩⟩ + simpa using hn + exists_of_eq {x y} h := by + obtain ⟨n, hn⟩ := exists_of_eq x y h + use ⟨r ^ n, n, rfl⟩ + +lemma of_associated {r r' : R} (h : Associated r r') [IsLocalization.Away r S] : + IsLocalization.Away r' S := by + obtain ⟨u, rfl⟩ := h + refine mk _ ?_ (fun s ↦ ?_) (fun a b hab ↦ ?_) + · simp [algebraMap_isUnit r, IsUnit.map _ u.isUnit] + · obtain ⟨n, a, hn⟩ := surj r s + use n, a * u ^ n + simp [mul_pow, ← mul_assoc, hn] + · obtain ⟨n, hn⟩ := exists_of_eq r hab + use n + rw [mul_pow, mul_comm (r ^ n), mul_assoc, mul_assoc, hn] + +/-- If `r` and `r'` are associated elements of `R`, an `R`-algebra `S` +is the localization of `R` away from `r` if and only of it is the localization of `R` away from +`r'`. -/ +lemma iff_of_associated {r r' : R} (h : Associated r r') : + IsLocalization.Away r S ↔ IsLocalization.Away r' S := + ⟨fun _ ↦ IsLocalization.Away.of_associated h, fun _ ↦ IsLocalization.Away.of_associated h.symm⟩ + variable {g : R →+* P} /-- Given `x : R`, a localization map `F : R →+* S` away from `x`, and a map of `CommSemiring`s @@ -146,6 +199,53 @@ lemma mapₐ_surjective_of_surjective {f : A →ₐ[R] B} (a : A) [Away a Aₚ] end Algebra +/-- Localizing the localization of `R` at `x` at the image of `y` is the same as localizing +`R` at `y * x`. See `IsLocalization.Away.mul'` for the `x * y` version. -/ +lemma mul (T : Type*) [CommSemiring T] [Algebra S T] + [Algebra R T] [IsScalarTower R S T] (x y : R) + [IsLocalization.Away x S] [IsLocalization.Away (algebraMap R S y) T] : + IsLocalization.Away (y * x) T := by + refine mk _ ?_ (fun z ↦ ?_) (fun a b h ↦ ?_) + · simp only [map_mul, IsUnit.mul_iff, IsScalarTower.algebraMap_apply R S T] + exact ⟨algebraMap_isUnit _, IsUnit.map _ (algebraMap_isUnit x)⟩ + · obtain ⟨m, p, hpq⟩ := surj (algebraMap R S y) z + obtain ⟨n, a, hab⟩ := surj x p + use m + n, a * x ^ m * y ^ n + simp only [mul_pow, pow_add, map_pow, map_mul, ← mul_assoc, hpq, + IsScalarTower.algebraMap_apply R S T, ← hab] + ring + · repeat rw [IsScalarTower.algebraMap_apply R S T] at h + obtain ⟨n, hn⟩ := exists_of_eq (algebraMap R S y) h + simp only [← map_pow, ← map_mul, ← map_mul] at hn + obtain ⟨m, hm⟩ := exists_of_eq x hn + use n + m + convert_to y ^ m * x ^ n * (x ^ m * (y ^ n * a)) = y ^ m * x ^ n * (x ^ m * (y ^ n * b)) + · ring + · ring + · rw [hm] + +/-- Localizing the localization of `R` at `x` at the image of `y` is the same as localizing +`R` at `x * y`. See `IsLocalization.Away.mul` for the `y * x` version. -/ +lemma mul' (T : Type*) [CommSemiring T] [Algebra S T] [Algebra R T] [IsScalarTower R S T] (x y : R) + [IsLocalization.Away x S] [IsLocalization.Away (algebraMap R S y) T] : + IsLocalization.Away (x * y) T := + mul_comm x y ▸ mul S T x y + +/-- If `S₁` is the localization of `R` away from `f` and `S₂` is the localization away from `g`, +then any localization `T` of `S₂` away from `f` is also a localization of `S₁` away from `g`. -/ +lemma commutes {R : Type*} [CommSemiring R] (S₁ S₂ T : Type*) [CommSemiring S₁] + [CommSemiring S₂] [CommSemiring T] [Algebra R S₁] [Algebra R S₂] [Algebra R T] [Algebra S₁ T] + [Algebra S₂ T] [IsScalarTower R S₁ T] [IsScalarTower R S₂ T] (x y : R) + [IsLocalization.Away x S₁] [IsLocalization.Away y S₂] + [IsLocalization.Away (algebraMap R S₂ x) T] : + IsLocalization.Away (algebraMap R S₁ y) T := by + haveI : IsLocalization (Algebra.algebraMapSubmonoid S₂ (Submonoid.powers x)) T := by + simp only [Algebra.algebraMapSubmonoid, Submonoid.map_powers] + infer_instance + convert IsLocalization.commutes S₁ S₂ T (Submonoid.powers x) (Submonoid.powers y) + ext x + simp [Algebra.algebraMapSubmonoid] + end Away end Away diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index b6a8a172936cf..ee3d1fbf1a3d4 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -828,6 +828,48 @@ theorem map_nonZeroDivisors_le [IsLocalization M S] : (nonZeroDivisors R).map (algebraMap R S) ≤ nonZeroDivisors S := Submonoid.map_le_iff_le_comap.mpr (nonZeroDivisors_le_comap M S) +/-- If `S₁` is the localization of `R` at `M₁` and `S₂` is the localization of +`R` at `M₂`, then every localization `T` of `S₂` at `M₁` is also a localization of +`S₁` at `M₂`, in other words `M₁⁻¹M₂⁻¹R` can be identified with `M₂⁻¹M₁⁻¹R`. -/ +lemma commutes (S₁ S₂ T : Type*) [CommSemiring S₁] + [CommSemiring S₂] [CommSemiring T] [Algebra R S₁] [Algebra R S₂] [Algebra R T] [Algebra S₁ T] + [Algebra S₂ T] [IsScalarTower R S₁ T] [IsScalarTower R S₂ T] (M₁ M₂ : Submonoid R) + [IsLocalization M₁ S₁] [IsLocalization M₂ S₂] + [IsLocalization (Algebra.algebraMapSubmonoid S₂ M₁) T] : + IsLocalization (Algebra.algebraMapSubmonoid S₁ M₂) T where + map_units' := by + rintro ⟨m, ⟨a, ha, rfl⟩⟩ + rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R S₂ T] + exact IsUnit.map _ (IsLocalization.map_units' ⟨a, ha⟩) + surj' a := by + obtain ⟨⟨y, -, m, hm, rfl⟩, hy⟩ := surj (M := Algebra.algebraMapSubmonoid S₂ M₁) a + rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R S₁ T] at hy + obtain ⟨⟨z, n, hn⟩, hz⟩ := IsLocalization.surj (M := M₂) y + have hunit : IsUnit (algebraMap R S₁ m) := map_units' ⟨m, hm⟩ + use ⟨algebraMap R S₁ z * hunit.unit⁻¹, ⟨algebraMap R S₁ n, n, hn, rfl⟩⟩ + rw [map_mul, ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R S₂ T] + conv_rhs => rw [← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply R S₂ T, ← hz, map_mul, ← hy] + convert_to _ = a * (algebraMap S₂ T) ((algebraMap R S₂) n) * + (algebraMap S₁ T) (((algebraMap R S₁) m) * hunit.unit⁻¹.val) + · rw [map_mul] + ring + simp + exists_of_eq {x y} hxy := by + obtain ⟨r, s, d, hr, hs⟩ := IsLocalization.surj₂ M₁ S₁ x y + apply_fun (· * algebraMap S₁ T (algebraMap R S₁ d)) at hxy + simp_rw [← map_mul, hr, hs, ← IsScalarTower.algebraMap_apply, + IsScalarTower.algebraMap_apply R S₂ T] at hxy + obtain ⟨⟨-, c, hmc, rfl⟩, hc⟩ := exists_of_eq (M := Algebra.algebraMapSubmonoid S₂ M₁) hxy + simp_rw [← map_mul] at hc + obtain ⟨a, ha⟩ := IsLocalization.exists_of_eq (M := M₂) hc + use ⟨algebraMap R S₁ a, a, a.property, rfl⟩ + apply (map_units S₁ d).mul_right_cancel + rw [mul_assoc, hr, mul_assoc, hs] + apply (map_units S₁ ⟨c, hmc⟩).mul_right_cancel + rw [← map_mul, ← map_mul, mul_assoc, mul_comm _ c, ha, map_mul, map_mul] + ring + end IsLocalization namespace Localization From e060e71a8717aaba9cd7fcf5df82ea51bc0e6d64 Mon Sep 17 00:00:00 2001 From: Jiazhen Xia Date: Sun, 29 Sep 2024 14:59:03 +0000 Subject: [PATCH 040/174] feat(Topology): define (relative) CW-complexes (#12502) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define a relative CW-complex as the colimit of an expanding sequence of subspaces `sk i` (called the $(i-1)$-skeleton) for `i ≥ 0`, where `sk 0` (i.e., the $(-1)$-skeleton) is an arbitrary topological space, and each `sk (n + 1)` (i.e., the $n$-skeleton) is obtained from `sk n` (i.e., the $(n-1)$-skeleton) by attaching `n`-disks. This is based on [David Wärn's suggestion on Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Do.20we.20have.20CW.20complexes.3F/near/231769080) and Chapter 10 of _[A Concise Course in Algebraic Topology](https://www.math.uchicago.edu/~may/CONCISE/ConciseRevised.pdf)_ by J. P. May. Co-authored-by: Elliot Dean Young Co-authored-by: Jiazhen Xia <6234949+jzxia@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Topology/CWComplex.lean | 97 +++++++++++++++++++ Mathlib/Topology/Category/TopCat/Basic.lean | 17 ++++ .../Category/TopCat/Limits/Basic.lean | 2 +- docs/references.bib | 12 +++ 5 files changed, 128 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Topology/CWComplex.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8ddcc87510e87..790efd9e50ebc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4503,6 +4503,7 @@ import Mathlib.Topology.Bornology.Basic import Mathlib.Topology.Bornology.BoundedOperation import Mathlib.Topology.Bornology.Constructions import Mathlib.Topology.Bornology.Hom +import Mathlib.Topology.CWComplex import Mathlib.Topology.Category.Born import Mathlib.Topology.Category.CompHaus.Basic import Mathlib.Topology.Category.CompHaus.EffectiveEpi diff --git a/Mathlib/Topology/CWComplex.lean b/Mathlib/Topology/CWComplex.lean new file mode 100644 index 0000000000000..d4bad250bb01d --- /dev/null +++ b/Mathlib/Topology/CWComplex.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2024 Elliot Dean Young and Jiazhen Xia. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiazhen Xia, Elliot Dean Young +-/ +import Mathlib.Topology.Category.TopCat.Limits.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Products +import Mathlib.CategoryTheory.Functor.OfSequence + +/-! +# CW-complexes + +This file defines (relative) CW-complexes. + +## Main definitions + +* `RelativeCWComplex`: A relative CW-complex is the colimit of an expanding sequence of subspaces + `sk i` (called the $(i-1)$-skeleton) for `i ≥ 0`, where `sk 0` (i.e., the $(-1)$-skeleton) is an + arbitrary topological space, and each `sk (n + 1)` (i.e., the $n$-skeleton) is obtained from + `sk n` (i.e., the $(n-1)$-skeleton) by attaching `n`-disks. + +* `CWComplex`: A CW-complex is a relative CW-complex whose `sk 0` (i.e., $(-1)$-skeleton) is empty. + +## References + +* [R. Fritsch and R. Piccinini, *Cellular Structures in Topology*][fritsch-piccinini1990] +* The definition of CW-complexes follows David Wärn's suggestion on + [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Do.20we.20have.20CW.20complexes.3F/near/231769080). +-/ + +open CategoryTheory TopCat + +universe u + +namespace RelativeCWComplex + +/-- The inclusion map from the `n`-sphere to the `(n + 1)`-disk -/ +def sphereInclusion (n : ℤ) : 𝕊 n ⟶ 𝔻 (n + 1) where + toFun := fun ⟨p, hp⟩ ↦ ⟨p, le_of_eq hp⟩ + continuous_toFun := ⟨fun t ⟨s, ⟨r, hro, hrs⟩, hst⟩ ↦ by + rw [isOpen_induced_iff, ← hst, ← hrs] + tauto⟩ + +/-- A type witnessing that `X'` is obtained from `X` by attaching generalized cells `f : S ⟶ D` -/ +structure AttachGeneralizedCells {S D : TopCat.{u}} (f : S ⟶ D) (X X' : TopCat.{u}) where + /-- The index type over the generalized cells -/ + cells : Type u + /-- An attaching map for each generalized cell -/ + attachMaps : cells → (S ⟶ X) + /-- `X'` is the pushout of `∐ S ⟶ X` and `∐ S ⟶ ∐ D`. -/ + iso_pushout : X' ≅ Limits.pushout (Limits.Sigma.desc attachMaps) (Limits.Sigma.map fun _ ↦ f) + +/-- A type witnessing that `X'` is obtained from `X` by attaching `(n + 1)`-disks -/ +def AttachCells (n : ℤ) := AttachGeneralizedCells (sphereInclusion n) + +end RelativeCWComplex + +/-- A relative CW-complex consists of an expanding sequence of subspaces `sk i` (called the +$(i-1)$-skeleton) for `i ≥ 0`, where `sk 0` (i.e., the $(-1)$-skeleton) is an arbitrary topological +space, and each `sk (n + 1)` (i.e., the `n`-skeleton) is obtained from `sk n` (i.e., the +$(n-1)$-skeleton) by attaching `n`-disks. -/ +structure RelativeCWComplex where + /-- The skeletons. Note: `sk i` is usually called the $(i-1)$-skeleton in the math literature. -/ + sk : ℕ → TopCat.{u} + /-- Each `sk (n + 1)` (i.e., the $n$-skeleton) is obtained from `sk n` + (i.e., the $(n-1)$-skeleton) by attaching `n`-disks. -/ + attachCells (n : ℕ) : RelativeCWComplex.AttachCells ((n : ℤ) - 1) (sk n) (sk (n + 1)) + +/-- A CW-complex is a relative CW-complex whose `sk 0` (i.e., $(-1)$-skeleton) is empty. -/ +structure CWComplex extends RelativeCWComplex.{u} where + /-- `sk 0` (i.e., the $(-1)$-skeleton) is empty. -/ + isEmpty_sk_zero : IsEmpty (sk 0) + +namespace RelativeCWComplex + +noncomputable section Topology + +/-- The inclusion map from `X` to `X'`, given that `X'` is obtained from `X` by attaching +`(n + 1)`-disks -/ +def AttachCells.inclusion {X X' : TopCat.{u}} {n : ℤ} (att : AttachCells n X X') : X ⟶ X' := + Limits.pushout.inl (Limits.Sigma.desc att.attachMaps) + (Limits.Sigma.map fun _ ↦ sphereInclusion n) ≫ att.iso_pushout.inv + +/-- The inclusion map from `sk n` (i.e., the $(n-1)$-skeleton) to `sk (n + 1)` (i.e., the +$n$-skeleton) of a relative CW-complex -/ +def skInclusion (X : RelativeCWComplex.{u}) (n : ℕ) : X.sk n ⟶ X.sk (n + 1) := + (X.attachCells n).inclusion + +/-- The topology on a relative CW-complex -/ +def toTopCat (X : RelativeCWComplex.{u}) : TopCat.{u} := + Limits.colimit (Functor.ofSequence X.skInclusion) + +instance : Coe RelativeCWComplex TopCat where coe X := toTopCat X + +end Topology + +end RelativeCWComplex diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 40459bf75dc71..3e3e458b83d31 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot, Kim Morrison, Mario Carneiro -/ import Mathlib.CategoryTheory.ConcreteCategory.BundledHom import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Analysis.InnerProductSpace.PiL2 /-! # Category instance for topological spaces @@ -187,4 +188,20 @@ theorem openEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ simp only [← Functor.map_comp] exact openEmbedding_iff_isIso_comp f g +/-- The `n`-sphere is the set of points in ℝⁿ⁺¹ whose norm equals `1`, +endowed with the subspace topology. -/ +noncomputable def sphere (n : ℤ) : TopCat.{u} := + TopCat.of <| ULift <| Metric.sphere (0 : EuclideanSpace ℝ <| Fin <| (n + 1).toNat) 1 + +/-- The `n`-disk is the set of points in ℝⁿ whose norm is at most `1`, +endowed with the subspace topology. -/ +noncomputable def disk (n : ℤ) : TopCat.{u} := + TopCat.of <| ULift <| Metric.closedBall (0 : EuclideanSpace ℝ <| Fin <| n.toNat) 1 + +/-- `𝕊 n` denotes the `n`-sphere. -/ +scoped prefix:arg "𝕊 " => sphere + +/-- `𝔻 n` denotes the `n`-disk. -/ +scoped prefix:arg "𝔻 " => disk + end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index 3fcf13740229c..6cfe839d9f554 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -184,7 +184,7 @@ instance forgetPreservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ Ty /-- The terminal object of `Top` is `PUnit`. -/ def isTerminalPUnit : IsTerminal (TopCat.of PUnit.{u + 1}) := haveI : ∀ X, Unique (X ⟶ TopCat.of PUnit.{u + 1}) := fun X => - ⟨⟨⟨fun _ => PUnit.unit, by continuity⟩⟩, fun f => by ext; aesop⟩ + ⟨⟨⟨fun _ => PUnit.unit, continuous_const⟩⟩, fun f => by ext; aesop⟩ Limits.IsTerminal.ofUnique _ /-- The terminal object of `Top` is `PUnit`. -/ diff --git a/docs/references.bib b/docs/references.bib index 65985063312cd..2f63bf3e561a3 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1396,6 +1396,18 @@ @Book{ friedmanscarr2005 zbl = {1080.46001} } +@Book{ fritsch-piccinini1990, + place = {Cambridge}, + series = {Cambridge Studies in Advanced Mathematics}, + title = {Cellular Structures in Topology}, + publisher = {Cambridge University Press}, + author = {Fritsch, Rudolf and Piccinini, Renzo}, + year = {1990}, + collection = {Cambridge Studies in Advanced Mathematics}, + url = {https://doi.org/10.1017/CBO9780511983948}, + doi = {10.1017/CBO9780511983948} +} + @Book{ fuchs1963, author = {Fuchs, L.}, title = {Partially ordered algebraic systems}, From 54358ec938bcf95978586464cbd1010833f6970a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 29 Sep 2024 14:59:05 +0000 Subject: [PATCH 041/174] feat(UniformEmbedding): add `UniformInducing.completeSpace_congr` (#17238) This is a version of `_root_.completeSpace_congr` that applies to uniform inducing quotient maps. Also add `ULift.completeSpace_iff` and `UniformInducing.completeSpace`. I'm going to use these lemmas to generalize some `CompleteSpace` instances, see #17244. --- .../Analysis/Normed/Algebra/Unitization.lean | 2 +- Mathlib/Analysis/Quaternion.lean | 2 +- .../Algebra/Module/FiniteDimension.lean | 6 ++-- .../UniformSpace/UniformEmbedding.lean | 35 ++++++++++++------- 4 files changed, 28 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/Normed/Algebra/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean index 90fcffffc0c67..8b5d489952f9a 100644 --- a/Mathlib/Analysis/Normed/Algebra/Unitization.lean +++ b/Mathlib/Analysis/Normed/Algebra/Unitization.lean @@ -216,7 +216,7 @@ theorem uniformEmbedding_addEquiv {𝕜} [NontriviallyNormedField 𝕜] : /-- `Unitization 𝕜 A` is complete whenever `𝕜` and `A` are also. -/ instance instCompleteSpace [CompleteSpace 𝕜] [CompleteSpace A] : CompleteSpace (Unitization 𝕜 A) := - (completeSpace_congr uniformEmbedding_addEquiv).mpr CompleteSpace.prod + uniformEquivProd.completeSpace_iff.2 .prod /-- Pull back the metric structure from `𝕜 × (A →L[𝕜] A)` to `Unitization 𝕜 A` using the algebra homomorphism `Unitization.splitMul 𝕜 A`, but replace the bornology and the uniformity so diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index 0991d3a5949c9..e342a1859bb60 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -197,7 +197,7 @@ theorem continuous_im : Continuous fun q : ℍ => q.im := by instance : CompleteSpace ℍ := haveI : UniformEmbedding linearIsometryEquivTuple.toLinearEquiv.toEquiv.symm := linearIsometryEquivTuple.toContinuousLinearEquiv.symm.uniformEmbedding - (completeSpace_congr this).1 (by infer_instance) + (completeSpace_congr this).1 inferInstance section infinite_sum diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 8b1c1ca19d66a..04cdc12e8032b 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -212,7 +212,7 @@ private theorem continuous_equivFun_basis_aux [T2Space E] {ι : Type v} [Fintype exact b.equivFun.symm.uniformEmbedding b.equivFun.symm.toLinearMap.continuous_on_pi this have : IsComplete (s : Set E) := - completeSpace_coe_iff_isComplete.1 ((completeSpace_congr U).1 (by infer_instance)) + completeSpace_coe_iff_isComplete.1 ((completeSpace_congr U).1 inferInstance) exact this.isClosed -- second step: any linear form is continuous, as its kernel is closed by the first step have H₂ : ∀ f : E →ₗ[𝕜] 𝕜, Continuous f := by @@ -490,8 +490,8 @@ variable (𝕜 E : Type*) [NontriviallyNormedField 𝕜] include 𝕜 in theorem FiniteDimensional.complete [FiniteDimensional 𝕜 E] : CompleteSpace E := by set e := ContinuousLinearEquiv.ofFinrankEq (@finrank_fin_fun 𝕜 _ _ (finrank 𝕜 E)).symm - have : UniformEmbedding e.toLinearEquiv.toEquiv.symm := e.symm.uniformEmbedding - exact (completeSpace_congr this).1 (by infer_instance) + have : UniformEmbedding e.toEquiv.symm := e.symm.uniformEmbedding + exact (completeSpace_congr this).1 inferInstance variable {𝕜 E} diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index b8882ff944db5..9029d071993e6 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -286,27 +286,36 @@ theorem completeSpace_iff_isComplete_range {f : α → β} (hf : UniformInducing CompleteSpace α ↔ IsComplete (range f) := by rw [completeSpace_iff_isComplete_univ, ← isComplete_image_iff hf, image_univ] +alias ⟨_, UniformInducing.completeSpace⟩ := completeSpace_iff_isComplete_range + theorem UniformInducing.isComplete_range [CompleteSpace α] {f : α → β} (hf : UniformInducing f) : IsComplete (range f) := (completeSpace_iff_isComplete_range hf).1 ‹_› +/-- If `f` is a surjective uniform inducing map, +then its domain is a complete space iff its codomain is a complete space. +See also `_root_.completeSpace_congr` for a version that assumes `f` to be an equivalence. -/ +theorem UniformInducing.completeSpace_congr {f : α → β} (hf : UniformInducing f) + (hsurj : f.Surjective) : CompleteSpace α ↔ CompleteSpace β := by + rw [completeSpace_iff_isComplete_range hf, hsurj.range_eq, completeSpace_iff_isComplete_univ] + theorem SeparationQuotient.completeSpace_iff : - CompleteSpace (SeparationQuotient α) ↔ CompleteSpace α := by - rw [completeSpace_iff_isComplete_univ, ← range_mk, - ← completeSpace_iff_isComplete_range uniformInducing_mk] + CompleteSpace (SeparationQuotient α) ↔ CompleteSpace α := + .symm <| uniformInducing_mk.completeSpace_congr surjective_mk instance SeparationQuotient.instCompleteSpace [CompleteSpace α] : CompleteSpace (SeparationQuotient α) := completeSpace_iff.2 ‹_› +/-- See also `UniformInducing.completeSpace_congr` +for a version that works for non-injective maps. -/ theorem completeSpace_congr {e : α ≃ β} (he : UniformEmbedding e) : - CompleteSpace α ↔ CompleteSpace β := by - rw [completeSpace_iff_isComplete_range he.toUniformInducing, e.range_eq_univ, - completeSpace_iff_isComplete_univ] + CompleteSpace α ↔ CompleteSpace β := + he.completeSpace_congr e.surjective -theorem completeSpace_coe_iff_isComplete {s : Set α} : CompleteSpace s ↔ IsComplete s := - (completeSpace_iff_isComplete_range uniformEmbedding_subtype_val.toUniformInducing).trans <| by - rw [Subtype.range_coe] +theorem completeSpace_coe_iff_isComplete {s : Set α} : CompleteSpace s ↔ IsComplete s := by + rw [completeSpace_iff_isComplete_range uniformEmbedding_subtype_val.toUniformInducing, + Subtype.range_coe] alias ⟨_, IsComplete.completeSpace_coe⟩ := completeSpace_coe_iff_isComplete @@ -314,10 +323,12 @@ theorem IsClosed.completeSpace_coe [CompleteSpace α] {s : Set α} (hs : IsClose CompleteSpace s := hs.isComplete.completeSpace_coe +theorem completeSpace_ulift_iff : CompleteSpace (ULift α) ↔ CompleteSpace α := + UniformInducing.completeSpace_congr ⟨rfl⟩ ULift.down_surjective + /-- The lift of a complete space to another universe is still complete. -/ -instance ULift.completeSpace [h : CompleteSpace α] : CompleteSpace (ULift α) := - haveI : UniformEmbedding (@Equiv.ulift α) := ⟨⟨rfl⟩, ULift.down_injective⟩ - (completeSpace_congr this).2 h +instance ULift.instCompleteSpace [CompleteSpace α] : CompleteSpace (ULift α) := + completeSpace_ulift_iff.2 ‹_› theorem completeSpace_extension {m : β → α} (hm : UniformInducing m) (dense : DenseRange m) (h : ∀ f : Filter β, Cauchy f → ∃ x : α, map m f ≤ 𝓝 x) : CompleteSpace α := From cea0638b3fffb0605c53cddbb166a4b918c74e82 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 29 Sep 2024 17:58:12 +0000 Subject: [PATCH 042/174] feat: `fin_omega`, a wrapper around `omega` to help with `Fin` arithmetic goals. (#16651) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is a hack, and that results in `omega` having to deal with a lot of case splitting, but it is usable. The simp set could probably be expanded a bit, and I've left a comment inviting such changes. This was inspired by #15817, and in particular replaces the new proof added in that PR with `by fin_omega`. Co-authored-by: Iván Renison <85908989+IvanRenison@users.noreply.github.com> Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Data/Fin/Basic.lean | 85 ++++++++++++++++++++++--------- Mathlib/Data/Int/DivMod.lean | 20 ++++++++ Mathlib/Tactic/Attr/Register.lean | 3 ++ 4 files changed, 85 insertions(+), 24 deletions(-) create mode 100644 Mathlib/Data/Int/DivMod.lean diff --git a/Mathlib.lean b/Mathlib.lean index 790efd9e50ebc..187e14e57636f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2251,6 +2251,7 @@ import Mathlib.Data.Int.Cast.Prod import Mathlib.Data.Int.CharZero import Mathlib.Data.Int.ConditionallyCompleteOrder import Mathlib.Data.Int.Defs +import Mathlib.Data.Int.DivMod import Mathlib.Data.Int.GCD import Mathlib.Data.Int.Interval import Mathlib.Data.Int.LeastGreatest diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 7763f1a2b0f45..d4c6eeb432941 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -5,9 +5,11 @@ Authors: Robert Y. Lewis, Keeley Hoek -/ import Mathlib.Algebra.NeZero import Mathlib.Data.Nat.Defs +import Mathlib.Data.Int.DivMod import Mathlib.Logic.Embedding.Basic import Mathlib.Logic.Equiv.Set import Mathlib.Tactic.Common +import Mathlib.Tactic.Attr.Register /-! # The finite type with `n` elements @@ -190,6 +192,7 @@ protected theorem heq_ext_iff {k l : ℕ} (h : k = l) {i : Fin k} {j : Fin l} : end coe + section Order /-! @@ -332,6 +335,58 @@ theorem one_lt_last [NeZero n] : 1 < last (n + 1) := by end Order +/-! ### Coercions to `ℤ` and the `fin_omega` tactic. -/ + +open Int + +theorem coe_int_sub_eq_ite {n : Nat} (u v : Fin n) : + ((u - v : Fin n) : Int) = if v ≤ u then (u - v : Int) else (u - v : Int) + n := by + rw [Fin.sub_def] + split + · rw [ofNat_emod, Int.emod_eq_sub_self_emod, Int.emod_eq_of_lt] <;> omega + · rw [ofNat_emod, Int.emod_eq_of_lt] <;> omega + +theorem coe_int_sub_eq_mod {n : Nat} (u v : Fin n) : + ((u - v : Fin n) : Int) = ((u : Int) - (v : Int)) % n := by + rw [coe_int_sub_eq_ite] + split + · rw [Int.emod_eq_of_lt] <;> omega + · rw [Int.emod_eq_add_self_emod, Int.emod_eq_of_lt] <;> omega + +theorem coe_int_add_eq_ite {n : Nat} (u v : Fin n) : + ((u + v : Fin n) : Int) = if (u + v : ℕ) < n then (u + v : Int) else (u + v : Int) - n := by + rw [Fin.add_def] + split + · rw [ofNat_emod, Int.emod_eq_of_lt] <;> omega + · rw [ofNat_emod, Int.emod_eq_sub_self_emod, Int.emod_eq_of_lt] <;> omega + +theorem coe_int_add_eq_mod {n : Nat} (u v : Fin n) : + ((u + v : Fin n) : Int) = ((u : Int) + (v : Int)) % n := by + rw [coe_int_add_eq_ite] + split + · rw [Int.emod_eq_of_lt] <;> omega + · rw [Int.emod_eq_sub_self_emod, Int.emod_eq_of_lt] <;> omega + +-- Write `a + b` as `if (a + b : ℕ) < n then (a + b : ℤ) else (a + b : ℤ) - n` and +-- similarly `a - b` as `if (b : ℕ) ≤ a then (a - b : ℤ) else (a - b : ℤ) + n`. +attribute [fin_omega] coe_int_sub_eq_ite coe_int_add_eq_ite + +-- Rewrite inequalities in `Fin` to inequalities in `ℕ` +attribute [fin_omega] Fin.lt_iff_val_lt_val Fin.le_iff_val_le_val + +-- Rewrite `1 : Fin (n + 2)` to `1 : ℤ` +attribute [fin_omega] val_one + +/-- +Preprocessor for `omega` to handle inequalities in `Fin`. +Note that this involves a lot of case splitting, so may be slow. +-/ +-- Further adjustment to the simp set can probably make this more powerful. +-- Please experiment and PR updates! +macro "fin_omega" : tactic => `(tactic| + { try simp only [fin_omega, ← Int.ofNat_lt, ← Int.ofNat_le] at * + omega }) + section Add /-! @@ -1439,9 +1494,8 @@ instance uniqueFinOne : Unique (Fin 1) where @[simp] theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a 0] -lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := - fin_two_eq_of_eq_zero_iff - (by simpa only [one_eq_zero_iff, succ.injEq, iff_false, reduceCtorEq] using hi) +lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := by + fin_omega @[simp] theorem coe_neg_one : ↑(-1 : Fin (n + 1)) = n := by @@ -1454,15 +1508,7 @@ theorem last_sub (i : Fin (n + 1)) : last n - i = Fin.rev i := Fin.ext <| by rw [coe_sub_iff_le.2 i.le_last, val_last, val_rev, Nat.succ_sub_succ_eq_sub] theorem add_one_le_of_lt {n : ℕ} {a b : Fin (n + 1)} (h : a < b) : a + 1 ≤ b := by - cases' a with a ha - cases' b with b hb - cases n - · simp only [Nat.zero_add, Nat.lt_one_iff] at ha hb - simp [ha, hb] - simp only [le_iff_val_le_val, val_add, lt_iff_val_lt_val, val_mk, val_one] at h ⊢ - rwa [Nat.mod_eq_of_lt, Nat.succ_le_iff] - rw [Nat.succ_lt_succ_iff] - exact h.trans_le (Nat.le_of_lt_succ hb) + cases n <;> fin_omega theorem exists_eq_add_of_le {n : ℕ} {a b : Fin n} (h : a ≤ b) : ∃ k ≤ b, b = a + k := by obtain ⟨k, hk⟩ : ∃ k : ℕ, (b : ℕ) = a + k := Nat.exists_eq_add_of_le h @@ -1473,15 +1519,10 @@ theorem exists_eq_add_of_le {n : ℕ} {a b : Fin n} (h : a ≤ b) : ∃ k ≤ b, theorem exists_eq_add_of_lt {n : ℕ} {a b : Fin (n + 1)} (h : a < b) : ∃ k < b, k + 1 ≤ b ∧ b = a + k + 1 := by cases n - · cases' a with a ha - cases' b with b hb - simp only [Nat.zero_add, Nat.lt_one_iff] at ha hb - simp [ha, hb] at h + · omega obtain ⟨k, hk⟩ : ∃ k : ℕ, (b : ℕ) = a + k + 1 := Nat.exists_eq_add_of_lt h have hkb : k < b := by omega - refine ⟨⟨k, hkb.trans b.is_lt⟩, hkb, ?_, ?_⟩ - · rw [Fin.le_iff_val_le_val, Fin.val_add_one] - split_ifs <;> simp [Nat.succ_le_iff, hkb] + refine ⟨⟨k, hkb.trans b.is_lt⟩, hkb, by fin_omega, ?_⟩ simp [Fin.ext_iff, Fin.val_add, ← hk, Nat.mod_eq_of_lt b.is_lt] lemma pos_of_ne_zero {n : ℕ} {a : Fin (n + 1)} (h : a ≠ 0) : @@ -1489,11 +1530,7 @@ lemma pos_of_ne_zero {n : ℕ} {a : Fin (n + 1)} (h : a ≠ 0) : Nat.pos_of_ne_zero (val_ne_of_ne h) lemma sub_succ_le_sub_of_le {n : ℕ} {u v : Fin (n + 2)} (h : u < v) : v - (u + 1) < v - u := by - have h' : u + 1 ≤ v := add_one_le_of_lt h - apply lt_def.mpr - simp only [sub_val_of_le h', sub_val_of_le (Fin.le_of_lt h)] - refine Nat.sub_lt_sub_left h (lt_def.mp ?_) - exact lt_add_one_iff.mpr (Fin.lt_of_lt_of_le h v.le_last) + fin_omega end AddGroup diff --git a/Mathlib/Data/Int/DivMod.lean b/Mathlib/Data/Int/DivMod.lean new file mode 100644 index 0000000000000..e2d8db47293a8 --- /dev/null +++ b/Mathlib/Data/Int/DivMod.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2024 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ + +/-! +# Basic lemmas about division and modulo for integers + +-/ + +namespace Int + +/-! ### `emod` -/ + +theorem emod_eq_sub_self_emod {a b : Int} : a % b = (a - b) % b := + (emod_sub_cancel a b).symm + +theorem emod_eq_add_self_emod {a b : Int} : a % b = (a + b) % b := + add_emod_self.symm diff --git a/Mathlib/Tactic/Attr/Register.lean b/Mathlib/Tactic/Attr/Register.lean index 98101a86367d7..cc80f10b4abbf 100644 --- a/Mathlib/Tactic/Attr/Register.lean +++ b/Mathlib/Tactic/Attr/Register.lean @@ -84,3 +84,6 @@ register_simp_attr nontriviality /-- A stub attribute for `is_poly`. -/ register_label_attr is_poly + +/-- A simp set for the `fin_omega` wrapper around `omega`. -/ +register_simp_attr fin_omega From 5a3902a416ddc0b9744888e8b0b246bcf18d4a76 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 29 Sep 2024 19:04:57 +0000 Subject: [PATCH 043/174] chore(Topology/Category): fix imports in TopCat.Basic (#17258) A new file `TopCat.Sphere` is added so as to remove the import to `Analysis.InnerProductSpace.PiL2` that was added to the file `Topology.Category.TopCat.Basic` in PR #12502. --- Mathlib.lean | 1 + Mathlib/Topology/CWComplex.lean | 15 +++++--- Mathlib/Topology/Category/TopCat/Basic.lean | 17 --------- Mathlib/Topology/Category/TopCat/Sphere.lean | 39 ++++++++++++++++++++ 4 files changed, 49 insertions(+), 23 deletions(-) create mode 100644 Mathlib/Topology/Category/TopCat/Sphere.lean diff --git a/Mathlib.lean b/Mathlib.lean index 187e14e57636f..277d2dffaca09 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4549,6 +4549,7 @@ import Mathlib.Topology.Category.TopCat.Limits.Products import Mathlib.Topology.Category.TopCat.Limits.Pullbacks import Mathlib.Topology.Category.TopCat.OpenNhds import Mathlib.Topology.Category.TopCat.Opens +import Mathlib.Topology.Category.TopCat.Sphere import Mathlib.Topology.Category.TopCat.Yoneda import Mathlib.Topology.Category.TopCommRingCat import Mathlib.Topology.Category.UniformSpace diff --git a/Mathlib/Topology/CWComplex.lean b/Mathlib/Topology/CWComplex.lean index d4bad250bb01d..8dae230580f73 100644 --- a/Mathlib/Topology/CWComplex.lean +++ b/Mathlib/Topology/CWComplex.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiazhen Xia, Elliot Dean Young -/ import Mathlib.Topology.Category.TopCat.Limits.Basic +import Mathlib.Topology.Category.TopCat.Sphere import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Functor.OfSequence @@ -34,7 +35,9 @@ universe u namespace RelativeCWComplex -/-- The inclusion map from the `n`-sphere to the `(n + 1)`-disk -/ +/-- The inclusion map from the `n`-sphere to the `(n + 1)`-disk. (For `n = -1`, this +involves the empty space `𝕊 (-1)`. This is the reason why `sphere` takes `n : ℤ` as +an input rather than `n : ℕ`.) -/ def sphereInclusion (n : ℤ) : 𝕊 n ⟶ 𝔻 (n + 1) where toFun := fun ⟨p, hp⟩ ↦ ⟨p, le_of_eq hp⟩ continuous_toFun := ⟨fun t ⟨s, ⟨r, hro, hrs⟩, hst⟩ ↦ by @@ -75,11 +78,11 @@ namespace RelativeCWComplex noncomputable section Topology -/-- The inclusion map from `X` to `X'`, given that `X'` is obtained from `X` by attaching -`(n + 1)`-disks -/ -def AttachCells.inclusion {X X' : TopCat.{u}} {n : ℤ} (att : AttachCells n X X') : X ⟶ X' := - Limits.pushout.inl (Limits.Sigma.desc att.attachMaps) - (Limits.Sigma.map fun _ ↦ sphereInclusion n) ≫ att.iso_pushout.inv +/-- The inclusion map from `X` to `X'`, when `X'` is obtained from `X` +by attaching generalized cells `f : S ⟶ D`. -/ +def AttachGeneralizedCells.inclusion {S D : TopCat.{u}} {f : S ⟶ D} {X X' : TopCat.{u}} + (att : AttachGeneralizedCells f X X') : X ⟶ X' := + Limits.pushout.inl _ _ ≫ att.iso_pushout.inv /-- The inclusion map from `sk n` (i.e., the $(n-1)$-skeleton) to `sk (n + 1)` (i.e., the $n$-skeleton) of a relative CW-complex -/ diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 3e3e458b83d31..40459bf75dc71 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -5,7 +5,6 @@ Authors: Patrick Massot, Kim Morrison, Mario Carneiro -/ import Mathlib.CategoryTheory.ConcreteCategory.BundledHom import Mathlib.Topology.ContinuousFunction.Basic -import Mathlib.Analysis.InnerProductSpace.PiL2 /-! # Category instance for topological spaces @@ -188,20 +187,4 @@ theorem openEmbedding_iff_isIso_comp' {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ simp only [← Functor.map_comp] exact openEmbedding_iff_isIso_comp f g -/-- The `n`-sphere is the set of points in ℝⁿ⁺¹ whose norm equals `1`, -endowed with the subspace topology. -/ -noncomputable def sphere (n : ℤ) : TopCat.{u} := - TopCat.of <| ULift <| Metric.sphere (0 : EuclideanSpace ℝ <| Fin <| (n + 1).toNat) 1 - -/-- The `n`-disk is the set of points in ℝⁿ whose norm is at most `1`, -endowed with the subspace topology. -/ -noncomputable def disk (n : ℤ) : TopCat.{u} := - TopCat.of <| ULift <| Metric.closedBall (0 : EuclideanSpace ℝ <| Fin <| n.toNat) 1 - -/-- `𝕊 n` denotes the `n`-sphere. -/ -scoped prefix:arg "𝕊 " => sphere - -/-- `𝔻 n` denotes the `n`-disk. -/ -scoped prefix:arg "𝔻 " => disk - end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Sphere.lean b/Mathlib/Topology/Category/TopCat/Sphere.lean new file mode 100644 index 0000000000000..b8f23e34fc643 --- /dev/null +++ b/Mathlib/Topology/Category/TopCat/Sphere.lean @@ -0,0 +1,39 @@ +/- +Copyright (c) 2024 Elliot Dean Young and Jiazhen Xia. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiazhen Xia, Elliot Dean Young +-/ + +import Mathlib.Analysis.InnerProductSpace.PiL2 +import Mathlib.Topology.Category.TopCat.Basic + +/-! +# Euclidean spheres + +This files defines the `n`-sphere `𝕊 n` and the `n`-disk `𝔻` as objects in `TopCat`. +The parameter `n` is in `ℤ` so as to facilitate the definition of +CW-complexes (see the file `Topology.CWComplex`). + +-/ + +universe u + +namespace TopCat + +/-- The `n`-sphere is the set of points in ℝⁿ⁺¹ whose norm equals `1`, +endowed with the subspace topology. -/ +noncomputable def sphere (n : ℤ) : TopCat.{u} := + TopCat.of <| ULift <| Metric.sphere (0 : EuclideanSpace ℝ <| Fin <| (n + 1).toNat) 1 + +/-- The `n`-disk is the set of points in ℝⁿ whose norm is at most `1`, +endowed with the subspace topology. -/ +noncomputable def disk (n : ℤ) : TopCat.{u} := + TopCat.of <| ULift <| Metric.closedBall (0 : EuclideanSpace ℝ <| Fin <| n.toNat) 1 + +/-- `𝕊 n` denotes the `n`-sphere. -/ +scoped prefix:arg "𝕊 " => sphere + +/-- `𝔻 n` denotes the `n`-disk. -/ +scoped prefix:arg "𝔻 " => disk + +end TopCat From 85e29307e7dd45b763f368e20254ae6d90b1a757 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Sun, 29 Sep 2024 20:16:33 +0000 Subject: [PATCH 044/174] feat(Combinatorics/Quiver/ReflQuiver): reflexive quivers (#16780) Reflexive quivers are an intermediate structure between categories and quivers, containing only an identity element. This PR constructs `ReflQuiver` as a class and `ReflQuiv` which is the category of reflexive quivers and reflexive prefunctors. Co-Authored-By: Emily Riehl and Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Co-authored-by: Mario Carneiro Co-authored-by: Emily Riehl Co-authored-by: Mario Carneiro --- Mathlib.lean | 2 + Mathlib/CategoryTheory/Category/Cat.lean | 8 + Mathlib/CategoryTheory/Category/Quiv.lean | 22 +- Mathlib/CategoryTheory/Category/ReflQuiv.lean | 251 ++++++++++++++++++ Mathlib/Combinatorics/Quiver/ReflQuiver.lean | 131 +++++++++ 5 files changed, 405 insertions(+), 9 deletions(-) create mode 100644 Mathlib/CategoryTheory/Category/ReflQuiv.lean create mode 100644 Mathlib/Combinatorics/Quiver/ReflQuiver.lean diff --git a/Mathlib.lean b/Mathlib.lean index 277d2dffaca09..9f3160338f29b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1445,6 +1445,7 @@ import Mathlib.CategoryTheory.Category.PartialFun import Mathlib.CategoryTheory.Category.Pointed import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.Category.Quiv +import Mathlib.CategoryTheory.Category.ReflQuiv import Mathlib.CategoryTheory.Category.RelCat import Mathlib.CategoryTheory.Category.TwoP import Mathlib.CategoryTheory.Category.ULift @@ -1960,6 +1961,7 @@ import Mathlib.Combinatorics.Quiver.ConnectedComponent import Mathlib.Combinatorics.Quiver.Covering import Mathlib.Combinatorics.Quiver.Path import Mathlib.Combinatorics.Quiver.Push +import Mathlib.Combinatorics.Quiver.ReflQuiver import Mathlib.Combinatorics.Quiver.SingleObj import Mathlib.Combinatorics.Quiver.Subquiver import Mathlib.Combinatorics.Quiver.Symmetric diff --git a/Mathlib/CategoryTheory/Category/Cat.lean b/Mathlib/CategoryTheory/Category/Cat.lean index 3d4788b016f0d..9838804baf16f 100644 --- a/Mathlib/CategoryTheory/Category/Cat.lean +++ b/Mathlib/CategoryTheory/Category/Cat.lean @@ -133,6 +133,14 @@ lemma associator_inv_app {B C D E : Cat} (F : B ⟶ C) (G : C ⟶ D) (H : D ⟶ (α_ F G H).inv.app X = eqToHom (by simp) := rfl +/-- The identity in the category of categories equals the identity functor.-/ +theorem id_eq_id (X : Cat) : 𝟙 X = 𝟭 X := rfl + +/-- Composition in the category of categories equals functor composition.-/ +theorem comp_eq_comp {X Y Z : Cat} (F : X ⟶ Y) (G : Y ⟶ Z) : F ≫ G = F ⋙ G := rfl + +@[simp] theorem of_α (C) [Category C] : (of C).α = C := rfl + /-- Functor that gets the set of objects of a category. It is not called `forget`, because it is not a faithful functor. -/ def objects : Cat.{v, u} ⥤ Type u where diff --git a/Mathlib/CategoryTheory/Category/Quiv.lean b/Mathlib/CategoryTheory/Category/Quiv.lean index bdb427995bad6..be83ad9b038c2 100644 --- a/Mathlib/CategoryTheory/Category/Quiv.lean +++ b/Mathlib/CategoryTheory/Category/Quiv.lean @@ -11,10 +11,8 @@ import Mathlib.CategoryTheory.PathCategory # The category of quivers The category of (bundled) quivers, and the free/forgetful adjunction between `Cat` and `Quiv`. - -/ - universe v u namespace CategoryTheory @@ -51,6 +49,12 @@ def forget : Cat.{v, u} ⥤ Quiv.{v, u} where obj C := Quiv.of C map F := F.toPrefunctor +/-- The identity in the category of quivers equals the identity prefunctor.-/ +theorem id_eq_id (X : Quiv) : 𝟙 X = 𝟭q X := rfl + +/-- Composition in the category of quivers equals prefunctor composition.-/ +theorem comp_eq_comp {X Y Z : Quiv} (F : X ⟶ Y) (G : Y ⟶ Z) : F ≫ G = F ⋙q G := rfl + end Quiv namespace Cat @@ -65,14 +69,14 @@ def free : Quiv.{v, u} ⥤ Cat.{max u v, u} where map_comp := fun f g => F.mapPath_comp f g } map_id V := by change (show Paths V ⥤ _ from _) = _ - ext; swap - · apply eq_conj_eqToHom + ext · rfl + · exact eq_conj_eqToHom _ map_comp {U _ _} F G := by change (show Paths U ⥤ _ from _) = _ - ext; swap - · apply eq_conj_eqToHom + ext · rfl + · exact eq_conj_eqToHom _ end Cat @@ -105,9 +109,9 @@ def adj : Cat.free ⊣ Quiv.forget := exact Category.id_comp _ } homEquiv_naturality_left_symm := fun {V _ _} f g => by change (show Paths V ⥤ _ from _) = _ - ext; swap - · apply eq_conj_eqToHom - · rfl } + ext + · rfl + · apply eq_conj_eqToHom } end Quiv diff --git a/Mathlib/CategoryTheory/Category/ReflQuiv.lean b/Mathlib/CategoryTheory/Category/ReflQuiv.lean new file mode 100644 index 0000000000000..f26afec53930b --- /dev/null +++ b/Mathlib/CategoryTheory/Category/ReflQuiv.lean @@ -0,0 +1,251 @@ +/- +Copyright (c) 2024 Mario Carneiro and Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl +-/ +import Mathlib.Combinatorics.Quiver.ReflQuiver +import Mathlib.CategoryTheory.Category.Cat +import Mathlib.CategoryTheory.Category.Quiv + +/-! +# The category of refl quivers + +The category `ReflQuiv` of (bundled) reflexive quivers, and the free/forgetful adjunction between +`Cat` and `ReflQuiv`. +-/ + +namespace CategoryTheory +universe v u + +/-- Category of refl quivers. -/ +@[nolint checkUnivs] +def ReflQuiv := + Bundled ReflQuiver.{v + 1, u} + +namespace ReflQuiv + +instance : CoeSort ReflQuiv (Type u) where coe := Bundled.α + +instance (C : ReflQuiv.{v, u}) : ReflQuiver.{v + 1, u} C := C.str + +/-- The underlying quiver of a reflexive quiver.-/ +def toQuiv (C : ReflQuiv.{v, u}) : Quiv.{v, u} := Quiv.of C.α + +/-- Construct a bundled `ReflQuiv` from the underlying type and the typeclass. -/ +def of (C : Type u) [ReflQuiver.{v + 1} C] : ReflQuiv.{v, u} := Bundled.of C + +instance : Inhabited ReflQuiv := ⟨ReflQuiv.of (Discrete default)⟩ + +@[simp] theorem of_val (C : Type u) [ReflQuiver C] : (ReflQuiv.of C) = C := rfl + +/-- Category structure on `ReflQuiv` -/ +instance category : LargeCategory.{max v u} ReflQuiv.{v, u} where + Hom C D := ReflPrefunctor C D + id C := ReflPrefunctor.id C + comp F G := ReflPrefunctor.comp F G + +theorem id_eq_id (X : ReflQuiv) : 𝟙 X = 𝟭rq X := rfl +theorem comp_eq_comp {X Y Z : ReflQuiv} (F : X ⟶ Y) (G : Y ⟶ Z) : F ≫ G = F ⋙rq G := rfl + +/-- The forgetful functor from categories to quivers. -/ +@[simps] +def forget : Cat.{v, u} ⥤ ReflQuiv.{v, u} where + obj C := ReflQuiv.of C + map F := F.toReflPrefunctor + +theorem forget_faithful {C D : Cat.{v, u}} (F G : C ⥤ D) + (hyp : forget.map F = forget.map G) : F = G := by + cases F; cases G; cases hyp; rfl + +theorem forget.Faithful : Functor.Faithful (forget) where + map_injective := fun hyp ↦ forget_faithful _ _ hyp + +/-- The forgetful functor from categories to quivers. -/ +@[simps] +def forgetToQuiv : ReflQuiv.{v, u} ⥤ Quiv.{v, u} where + obj V := Quiv.of V + map F := F.toPrefunctor + +theorem forgetToQuiv_faithful {V W : ReflQuiv} (F G : V ⥤rq W) + (hyp : forgetToQuiv.map F = forgetToQuiv.map G) : F = G := by + cases F; cases G; cases hyp; rfl + +theorem forgetToQuiv.Faithful : Functor.Faithful (forgetToQuiv) where + map_injective := fun hyp ↦ forgetToQuiv_faithful _ _ hyp + +theorem forget_forgetToQuiv : forget ⋙ forgetToQuiv = Quiv.forget := rfl + +end ReflQuiv + +namespace ReflPrefunctor + +/-- A refl prefunctor can be promoted to a functor if it respects composition.-/ +def toFunctor {C D : Cat} (F : (ReflQuiv.of C) ⟶ (ReflQuiv.of D)) + (hyp : ∀ {X Y Z : ↑C} (f : X ⟶ Y) (g : Y ⟶ Z), + F.map (CategoryStruct.comp (obj := C) f g) = + CategoryStruct.comp (obj := D) (F.map f) (F.map g)) : C ⥤ D where + obj := F.obj + map := F.map + map_id := F.map_id + map_comp := hyp + +end ReflPrefunctor + +namespace Cat + +/-- The hom relation that identifies the specified reflexivity arrows with the nil paths.-/ +inductive FreeReflRel {V} [ReflQuiver V] : (X Y : Paths V) → (f g : X ⟶ Y) → Prop + | mk {X : V} : FreeReflRel X X (Quiver.Hom.toPath (𝟙rq X)) .nil + +/-- A reflexive quiver generates a free category, defined as as quotient of the free category +on its underlying quiver (called the "path category") by the hom relation that uses the specified +reflexivity arrows as the identity arrows. -/ +def FreeRefl (V) [ReflQuiver V] := + Quotient (C := Cat.free.obj (Quiv.of V)) (FreeReflRel (V := V)) + +instance (V) [ReflQuiver V] : Category (FreeRefl V) := + inferInstanceAs (Category (Quotient _)) + +/-- The quotient functor associated to a quotient category defines a natural map from the free +category on the underlying quiver of a refl quiver to the free category on the reflexive quiver.-/ +def FreeRefl.quotientFunctor (V) [ReflQuiver V] : Cat.free.obj (Quiv.of V) ⥤ FreeRefl V := + Quotient.functor (C := Cat.free.obj (Quiv.of V)) (FreeReflRel (V := V)) + +/-- This is a specialization of `Quotient.lift_unique'` rather than `Quotient.lift_unique`, hence +the prime in the name.-/ +theorem FreeRefl.lift_unique' {V} [ReflQuiver V] {D} [Category D] (F₁ F₂ : FreeRefl V ⥤ D) + (h : quotientFunctor V ⋙ F₁ = quotientFunctor V ⋙ F₂) : + F₁ = F₂ := + Quotient.lift_unique' (C := Cat.free.obj (Quiv.of V)) (FreeReflRel (V := V)) _ _ h + +/-- The functor sending a reflexive quiver to the free category it generates, a quotient of +its path category.-/ +@[simps!] +def freeRefl : ReflQuiv.{v, u} ⥤ Cat.{max u v, u} where + obj V := Cat.of (FreeRefl V) + map f := Quotient.lift _ ((by exact Cat.free.map f.toPrefunctor) ⋙ FreeRefl.quotientFunctor _) + (fun X Y f g hfg => by + apply Quotient.sound + cases hfg + simp [ReflPrefunctor.map_id] + constructor) + map_id X := by + dsimp + refine (Quotient.lift_unique _ _ _ _ ((Functor.comp_id _).trans <| + (Functor.id_comp _).symm.trans ?_)).symm + congr 1 + exact (free.map_id X.toQuiv).symm + map_comp {X Y Z} f g := by + dsimp + apply (Quotient.lift_unique _ _ _ _ _).symm + have : free.map (f ≫ g).toPrefunctor = + free.map (X := X.toQuiv) (Y := Y.toQuiv) f.toPrefunctor ⋙ + free.map (X := Y.toQuiv) (Y := Z.toQuiv) g.toPrefunctor := by + show _ = _ ≫ _ + rw [← Functor.map_comp]; rfl + rw [this, Functor.assoc] + show _ ⋙ _ ⋙ _ = _ + rw [← Functor.assoc, Quotient.lift_spec, Functor.assoc, FreeRefl.quotientFunctor, + Quotient.lift_spec] + +theorem freeRefl_naturality {X Y} [ReflQuiver X] [ReflQuiver Y] (f : X ⥤rq Y) : + free.map (X := Quiv.of X) (Y := Quiv.of Y) f.toPrefunctor ⋙ + FreeRefl.quotientFunctor ↑Y = + FreeRefl.quotientFunctor ↑X ⋙ freeRefl.map (X := ReflQuiv.of X) (Y := ReflQuiv.of Y) f := by + simp only [free_obj, FreeRefl.quotientFunctor, freeRefl, ReflQuiv.of_val] + rw [Quotient.lift_spec] + +/-- We will make use of the natural quotient map from the free category on the underlying +quiver of a refl quiver to the free category on the reflexive quiver.-/ +def freeReflNatTrans : ReflQuiv.forgetToQuiv ⋙ Cat.free ⟶ freeRefl where + app V := FreeRefl.quotientFunctor V + naturality _ _ f := freeRefl_naturality f + +end Cat + +namespace ReflQuiv +open Category Functor + +/-- The unit components are defined as the composite of the corresponding unit component for the +adjunction between categories and quivers with the map underlying the quotient functor.-/ +@[simps! toPrefunctor obj map] +def adj.unit.app (V : ReflQuiv.{max u v, u}) : V ⥤rq forget.obj (Cat.freeRefl.obj V) where + toPrefunctor := Quiv.adj.unit.app (V.toQuiv) ⋙q + Quiv.forget.map (Cat.FreeRefl.quotientFunctor V) + map_id := fun _ => Quotient.sound _ ⟨⟩ + +/-- This is used in the proof of both triangle equalities.-/ +theorem adj.unit.component_eq (V : ReflQuiv.{max u v, u}) : + forgetToQuiv.map (adj.unit.app V) = Quiv.adj.unit.app (V.toQuiv) ≫ + Quiv.forget.map (Y := Cat.of _) (Cat.FreeRefl.quotientFunctor V) := rfl + +/-- The counit components are defined using the universal property of the quotient +from the corresponding counit component for the adjunction between categories and quivers.-/ +@[simps!] +def adj.counit.app (C : Cat) : Cat.freeRefl.obj (forget.obj C) ⥤ C := by + fapply Quotient.lift + · exact Quiv.adj.counit.app C + · intro x y f g rel + cases rel + unfold Quiv.adj + simp only [Adjunction.mkOfHomEquiv_counit_app, Equiv.coe_fn_symm_mk, + Quiv.lift_map, Prefunctor.mapPath_toPath, composePath_toPath] + rfl + +/-- The counit of `ReflQuiv.adj` is closely related to the counit of `Quiv.adj`.-/ +@[simp] +theorem adj.counit.component_eq (C : Cat) : + Cat.FreeRefl.quotientFunctor C ⋙ adj.counit.app C = + Quiv.adj.counit.app C := rfl + +/-- The counit of `ReflQuiv.adj` is closely related to the counit of `Quiv.adj`. For ease of use, +we introduce primed version for unbundled categories.-/ +@[simp] +theorem adj.counit.component_eq' (C) [Category C] : + Cat.FreeRefl.quotientFunctor C ⋙ adj.counit.app (Cat.of C) = + Quiv.adj.counit.app (Cat.of C) := rfl + +/-- +The adjunction between forming the free category on a reflexive quiver, and forgetting a category +to a reflexive quiver. +-/ +nonrec def adj : Cat.freeRefl.{max u v, u} ⊣ ReflQuiv.forget := + Adjunction.mkOfUnitCounit { + unit := { + app := adj.unit.app + naturality := fun V W f ↦ by exact rfl + } + counit := { + app := adj.counit.app + naturality := fun C D F ↦ Quotient.lift_unique' _ _ _ (Quiv.adj.counit.naturality F) + } + left_triangle := by + ext V + apply Cat.FreeRefl.lift_unique' + simp only [id_obj, Cat.free_obj, comp_obj, Cat.freeRefl_obj_α, NatTrans.comp_app, + forget_obj, whiskerRight_app, associator_hom_app, whiskerLeft_app, id_comp, + NatTrans.id_app'] + rw [Cat.id_eq_id, Cat.comp_eq_comp] + simp only [Cat.freeRefl_obj_α, Functor.comp_id] + rw [← Functor.assoc, ← Cat.freeRefl_naturality, Functor.assoc] + dsimp [Cat.freeRefl] + rw [adj.counit.component_eq' (Cat.FreeRefl V)] + conv => + enter [1, 1, 2] + apply (Quiv.comp_eq_comp (X := Quiv.of _) (Y := Quiv.of _) (Z := Quiv.of _) ..).symm + rw [Cat.free.map_comp] + show (_ ⋙ ((Quiv.forget ⋙ Cat.free).map (X := Cat.of _) (Y := Cat.of _) + (Cat.FreeRefl.quotientFunctor V))) ⋙ _ = _ + rw [Functor.assoc, ← Cat.comp_eq_comp] + conv => enter [1, 2]; apply Quiv.adj.counit.naturality + rw [Cat.comp_eq_comp, ← Functor.assoc, ← Cat.comp_eq_comp] + conv => enter [1, 1]; apply Quiv.adj.left_triangle_components V.toQuiv + exact Functor.id_comp _ + right_triangle := by + ext C + exact forgetToQuiv_faithful _ _ (Quiv.adj.right_triangle_components C) + } + +end ReflQuiv + +end CategoryTheory diff --git a/Mathlib/Combinatorics/Quiver/ReflQuiver.lean b/Mathlib/Combinatorics/Quiver/ReflQuiver.lean new file mode 100644 index 0000000000000..b0f9c85ac7250 --- /dev/null +++ b/Mathlib/Combinatorics/Quiver/ReflQuiver.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Mario Carneiro and Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl +-/ +import Mathlib.Data.Set.Function +import Mathlib.CategoryTheory.Category.Cat + +/-! +# Reflexive Quivers + +This module defines reflexive quivers. A reflexive quiver, or "refl quiver" for short, extends +a quiver with a specified endoarrow on each term in its type of objects. + +We also introduce morphisms between reflexive quivers, called reflexive prefunctors or "refl +prefunctors" for short. + +Note: Currently Category does not extend ReflQuiver, although it could. (TODO: do this) +-/ +namespace CategoryTheory +universe v v₁ v₂ u u₁ u₂ + +/-- A reflexive quiver extends a quiver with a specified arrow `id X : X ⟶ X` for each `X` in its +type of objects. We denote these arrows by `id` since categories can be understood as an extension +of refl quivers. +-/ +class ReflQuiver (obj : Type u) extends Quiver.{v} obj : Type max u v where + /-- The identity morphism on an object. -/ + id : ∀ X : obj, Hom X X + +/-- Notation for the identity morphism in a category. -/ +scoped notation "𝟙rq" => ReflQuiver.id -- type as \b1 + +instance catToReflQuiver {C : Type u} [inst : Category.{v} C] : ReflQuiver.{v+1, u} C := + { inst with } + +@[simp] theorem ReflQuiver.id_eq_id {C : Type*} [Category C] (X : C) : 𝟙rq X = 𝟙 X := rfl + +/-- A morphism of reflexive quivers called a `ReflPrefunctor`. -/ +structure ReflPrefunctor (V : Type u₁) [ReflQuiver.{v₁} V] (W : Type u₂) [ReflQuiver.{v₂} W] + extends Prefunctor V W where + /-- A functor preserves identity morphisms. -/ + map_id : ∀ X : V, map (𝟙rq X) = 𝟙rq (obj X) := by aesop_cat + +namespace ReflPrefunctor + +-- These lemmas can not be `@[simp]` because after `whnfR` they have a variable on the LHS. +-- Nevertheless they are sometimes useful when building functors. +lemma mk_obj {V W : Type*} [ReflQuiver V] [ReflQuiver W] {obj : V → W} {map} {X : V} : + (Prefunctor.mk obj map).obj X = obj X := rfl + +lemma mk_map {V W : Type*} [ReflQuiver V] [ReflQuiver W] {obj : V → W} {map} {X Y : V} {f : X ⟶ Y} : + (Prefunctor.mk obj map).map f = map f := rfl + +/-- Proving equality between reflexive prefunctors. This isn't an extensionality lemma, + because usually you don't really want to do this. -/ +theorem ext {V : Type u} [ReflQuiver.{v₁} V] {W : Type u₂} [ReflQuiver.{v₂} W] + {F G : ReflPrefunctor V W} + (h_obj : ∀ X, F.obj X = G.obj X) + (h_map : ∀ (X Y : V) (f : X ⟶ Y), + F.map f = Eq.recOn (h_obj Y).symm (Eq.recOn (h_obj X).symm (G.map f))) : F = G := by + obtain ⟨⟨F_obj⟩⟩ := F + obtain ⟨⟨G_obj⟩⟩ := G + obtain rfl : F_obj = G_obj := (Set.eqOn_univ F_obj G_obj).mp fun _ _ ↦ h_obj _ + congr + funext X Y f + simpa using h_map X Y f + +/-- The identity morphism between reflexive quivers. -/ +@[simps!] +def id (V : Type*) [ReflQuiver V] : ReflPrefunctor V V where + __ := Prefunctor.id _ + map_id _ := rfl + +instance (V : Type*) [ReflQuiver V] : Inhabited (ReflPrefunctor V V) := + ⟨id V⟩ + +/-- Composition of morphisms between reflexive quivers. -/ +@[simps!] +def comp {U : Type*} [ReflQuiver U] {V : Type*} [ReflQuiver V] {W : Type*} [ReflQuiver W] + (F : ReflPrefunctor U V) (G : ReflPrefunctor V W) : ReflPrefunctor U W where + __ := F.toPrefunctor.comp G.toPrefunctor + map_id _ := by simp [F.map_id, G.map_id] + +@[simp] +theorem comp_id {U V : Type*} [ReflQuiver U] [ReflQuiver V] (F : ReflPrefunctor U V) : + F.comp (id _) = F := rfl + +@[simp] +theorem id_comp {U V : Type*} [ReflQuiver U] [ReflQuiver V] (F : ReflPrefunctor U V) : + (id _).comp F = F := rfl + +@[simp] +theorem comp_assoc {U V W Z : Type*} [ReflQuiver U] [ReflQuiver V] [ReflQuiver W] [ReflQuiver Z] + (F : ReflPrefunctor U V) (G : ReflPrefunctor V W) (H : ReflPrefunctor W Z) : + (F.comp G).comp H = F.comp (G.comp H) := rfl + +/-- Notation for a prefunctor between reflexive quivers. -/ +infixl:50 " ⥤rq " => ReflPrefunctor + +/-- Notation for composition of reflexive prefunctors. -/ +infixl:60 " ⋙rq " => ReflPrefunctor.comp + +/-- Notation for the identity prefunctor on a reflexive quiver. -/ +notation "𝟭rq" => id + +theorem congr_map {U V : Type*} [Quiver U] [Quiver V] (F : U ⥤q V) {X Y : U} {f g : X ⟶ Y} + (h : f = g) : F.map f = F.map g := congrArg F.map h + +end ReflPrefunctor + +/-- A functor has an underlying refl prefunctor.-/ +def Functor.toReflPrefunctor {C D} [Category C] [Category D] (F : C ⥤ D) : C ⥤rq D := { F with } + +@[simp] +theorem Functor.toReflPrefunctor_toPrefunctor {C D : Cat} (F : C ⥤ D) : + (Functor.toReflPrefunctor F).toPrefunctor = F.toPrefunctor := rfl + +namespace ReflQuiver +open Opposite + +/-- `Vᵒᵖ` reverses the direction of all arrows of `V`. -/ +instance opposite {V} [ReflQuiver V] : ReflQuiver Vᵒᵖ where + id X := op (𝟙rq X.unop) + +instance discreteReflQuiver (V : Type u) : ReflQuiver.{u+1} (Discrete V) := + { discreteCategory V with } + +end ReflQuiver + +end CategoryTheory From d4f0e62ae56833e71eae899a03e89cc561932774 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Sep 2024 20:51:41 +0000 Subject: [PATCH 045/174] =?UTF-8?q?feat:=20`x=20=E2=88=88=20toZModSubmodul?= =?UTF-8?q?e=20n=20S=20=E2=86=94=20x=20=E2=88=88=20S`=20(#17210)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanAPAP --- Mathlib/Data/ZMod/Module.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/ZMod/Module.lean b/Mathlib/Data/ZMod/Module.lean index 64cb1d86b6160..c99250ffefff5 100644 --- a/Mathlib/Data/ZMod/Module.lean +++ b/Mathlib/Data/ZMod/Module.lean @@ -92,9 +92,8 @@ theorem toZModSubmodule_symm : ⇑((toZModSubmodule n).symm : _ ≃o AddSubgroup M) = Submodule.toAddSubgroup := rfl -@[simp] -theorem coe_toZModSubmodule (S : AddSubgroup M) : (toZModSubmodule n S : Set M) = S := - rfl +@[simp] lemma coe_toZModSubmodule (S : AddSubgroup M) : (toZModSubmodule n S : Set M) = S := rfl +@[simp] lemma mem_toZModSubmodule {S : AddSubgroup M} : x ∈ toZModSubmodule n S ↔ x ∈ S := .rfl @[simp] theorem toZModSubmodule_toAddSubgroup (S : AddSubgroup M) : From 301f54abf16c4d7706fb65a94ece16e384d6622e Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sun, 29 Sep 2024 21:20:19 +0000 Subject: [PATCH 046/174] feat(AlgebraicGeometry/EllipticCurve/Weierstrass): add `j_eq_zero[_iff][']` (#16974) which state the equivalency of `j = 0` and `c4 ^ 3 = 0` (and `c4 = 0` if the ring is reduced). Also added `IsReduced.pow_(eq|ne)_zero[_iff]`, which generalizes the `pow_(eq|ne)_zero[_iff]` assuming `NoZeroDivisors`. --- .../EllipticCurve/Weierstrass.lean | 10 ++++++++ Mathlib/RingTheory/Nilpotent/Defs.lean | 23 +++++++++++++++++++ 2 files changed, 33 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index d07fdd996d68d..a1e64d399e356 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -554,6 +554,16 @@ variable (E : EllipticCurve R) def j : R := E.Δ'⁻¹ * E.c₄ ^ 3 +/-- A variant of `EllipticCurve.j_eq_zero_iff` without assuming the ring being reduced. -/ +lemma j_eq_zero_iff' : E.j = 0 ↔ E.c₄ ^ 3 = 0 := by + rw [j, Units.mul_right_eq_zero] + +lemma j_eq_zero (h : E.c₄ = 0) : E.j = 0 := by + rw [j_eq_zero_iff', h, zero_pow three_ne_zero] + +lemma j_eq_zero_iff [IsReduced R] : E.j = 0 ↔ E.c₄ = 0 := by + rw [j_eq_zero_iff', IsReduced.pow_eq_zero_iff three_ne_zero] + lemma twoTorsionPolynomial_disc_ne_zero [Nontrivial R] [Invertible (2 : R)] : E.twoTorsionPolynomial.disc ≠ 0 := E.toWeierstrassCurve.twoTorsionPolynomial_disc_ne_zero <| E.coe_Δ' ▸ E.Δ'.isUnit diff --git a/Mathlib/RingTheory/Nilpotent/Defs.lean b/Mathlib/RingTheory/Nilpotent/Defs.lean index 0404b768832fa..cbdde6aab35bc 100644 --- a/Mathlib/RingTheory/Nilpotent/Defs.lean +++ b/Mathlib/RingTheory/Nilpotent/Defs.lean @@ -168,6 +168,29 @@ class IsReduced (R : Type*) [Zero R] [Pow R ℕ] : Prop where /-- A reduced structure has no nonzero nilpotent elements. -/ eq_zero : ∀ x : R, IsNilpotent x → x = 0 +namespace IsReduced + +theorem pow_eq_zero [Zero R] [Pow R ℕ] [IsReduced R] {n : ℕ} (h : x ^ n = 0) : + x = 0 := IsReduced.eq_zero x ⟨n, h⟩ + +@[simp] +theorem pow_eq_zero_iff [MonoidWithZero R] [IsReduced R] {n : ℕ} (hn : n ≠ 0) : + x ^ n = 0 ↔ x = 0 := ⟨pow_eq_zero, fun h ↦ h.symm ▸ zero_pow hn⟩ + +theorem pow_ne_zero_iff [MonoidWithZero R] [IsReduced R] {n : ℕ} (hn : n ≠ 0) : + x ^ n ≠ 0 ↔ x ≠ 0 := not_congr (pow_eq_zero_iff hn) + +theorem pow_ne_zero [Zero R] [Pow R ℕ] [IsReduced R] (n : ℕ) (h : x ≠ 0) : + x ^ n ≠ 0 := fun H ↦ h (pow_eq_zero H) + +/-- A variant of `IsReduced.pow_eq_zero_iff` assuming `R` is not trivial. -/ +@[simp] +theorem pow_eq_zero_iff' [MonoidWithZero R] [IsReduced R] [Nontrivial R] {n : ℕ} : + x ^ n = 0 ↔ x = 0 ∧ n ≠ 0 := by + cases n <;> simp + +end IsReduced + instance (priority := 900) isReduced_of_noZeroDivisors [MonoidWithZero R] [NoZeroDivisors R] : IsReduced R := ⟨fun _ ⟨_, hn⟩ => pow_eq_zero hn⟩ From 9f502c28baa1e9a8197bd0027eb6bfedc9e7fd3c Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 30 Sep 2024 00:42:10 +0000 Subject: [PATCH 047/174] chore: remove some unused variables (#17261) Inspired by #17178 and found by a linter. --- Mathlib/Algebra/Field/Basic.lean | 2 +- Mathlib/Algebra/Group/Defs.lean | 4 ++-- Mathlib/Algebra/Group/Equiv/Basic.lean | 1 - Mathlib/Algebra/Group/Semiconj/Units.lean | 2 +- Mathlib/Algebra/GroupWithZero/Basic.lean | 4 ++-- Mathlib/Algebra/GroupWithZero/Commute.lean | 4 ++-- Mathlib/Algebra/GroupWithZero/InjSurj.lean | 2 +- Mathlib/Algebra/GroupWithZero/Semiconj.lean | 2 +- Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean | 1 - Mathlib/Algebra/Module/BigOperators.lean | 2 +- Mathlib/Algebra/Polynomial/Induction.lean | 5 ++--- Mathlib/Algebra/Ring/Basic.lean | 9 +-------- Mathlib/Algebra/Ring/Commute.lean | 4 ++-- Mathlib/Algebra/Ring/Hom/Defs.lean | 4 +--- Mathlib/Algebra/Ring/Semiconj.lean | 6 +++--- Mathlib/Algebra/Ring/Units.lean | 2 +- Mathlib/Control/Basic.lean | 2 -- Mathlib/Control/Lawful.lean | 2 +- Mathlib/Data/Subtype.lean | 2 +- Mathlib/Lean/PrettyPrinter/Delaborator.lean | 7 ------- Mathlib/Logic/Basic.lean | 7 +++---- Mathlib/Logic/Equiv/Basic.lean | 2 -- Mathlib/Logic/Function/Conjugate.lean | 2 +- Mathlib/Logic/Relator.lean | 4 ++-- Mathlib/Logic/Unique.lean | 1 - Mathlib/Order/Basic.lean | 2 +- Mathlib/Order/Defs.lean | 1 - Mathlib/Order/Interval/Finset/Basic.lean | 8 ++++---- Mathlib/Order/Max.lean | 2 +- Mathlib/Order/SuccPred/Basic.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/Integral.lean | 4 ++-- Mathlib/Topology/AlexandrovDiscrete.lean | 2 -- Mathlib/Topology/MetricSpace/Infsep.lean | 6 +++--- Mathlib/Topology/MetricSpace/ShrinkingLemma.lean | 2 +- 34 files changed, 42 insertions(+), 70 deletions(-) diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index 570a5278b9b56..fe9961bb0230a 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -167,7 +167,7 @@ end DivisionRing section Semifield -variable [Semifield K] {a b c d : K} +variable [Semifield K] {a b d : K} theorem div_add_div (a : K) (c : K) (hb : b ≠ 0) (hd : d ≠ 0) : a / b + c / d = (a * d + b * c) / (b * d) := diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 810e4ced7a7a0..cd436c8a8fdfc 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -552,7 +552,7 @@ instance AddMonoid.toNatSMul {M : Type*} [AddMonoid M] : SMul ℕ M := attribute [to_additive existing toNatSMul] Monoid.toNatPow section Monoid -variable {M : Type*} [Monoid M] {a b c : M} {m n : ℕ} +variable {M : Type*} [Monoid M] {a b c : M} @[to_additive (attr := simp) nsmul_eq_smul] theorem npow_eq_pow (n : ℕ) (x : M) : Monoid.npow n x = x ^ n := @@ -879,7 +879,7 @@ theorem exists_zpow_surjective (G : Type*) [Pow G ℤ] [IsCyclic G] : section DivInvMonoid -variable [DivInvMonoid G] {a b : G} +variable [DivInvMonoid G] @[to_additive (attr := simp) zsmul_eq_smul] theorem zpow_eq_pow (n : ℤ) (x : G) : DivInvMonoid.zpow n x = x ^ n := diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 242b93f9ace47..ad1eee7b26f9c 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -130,7 +130,6 @@ instance (priority := 100) instMonoidHomClass _ = e (EquivLike.inv e (1 : N)) := by rw [← map_mul, one_mul] _ = 1 := EquivLike.right_inv e 1 } -variable [EquivLike F α β] variable {F} @[to_additive (attr := simp)] diff --git a/Mathlib/Algebra/Group/Semiconj/Units.lean b/Mathlib/Algebra/Group/Semiconj/Units.lean index b842e2c11bc3b..5a31a5e5130e9 100644 --- a/Mathlib/Algebra/Group/Semiconj/Units.lean +++ b/Mathlib/Algebra/Group/Semiconj/Units.lean @@ -32,7 +32,7 @@ assert_not_exists DenselyOrdered open scoped Int -variable {M G : Type*} +variable {M : Type*} namespace SemiconjBy diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 8b4b08535c105..26f0af2649201 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -291,7 +291,7 @@ end GroupWithZero section GroupWithZero -variable [GroupWithZero G₀] {a b c : G₀} +variable [GroupWithZero G₀] {a : G₀} @[simp] theorem zero_div (a : G₀) : 0 / a = 0 := by rw [div_eq_mul_inv, zero_mul] @@ -421,7 +421,7 @@ end GroupWithZero section CommGroupWithZero -variable [CommGroupWithZero G₀] {a b c d : G₀} +variable [CommGroupWithZero G₀] theorem div_mul_eq_mul_div₀ (a b c : G₀) : a / c * b = a * b / c := by simp_rw [div_eq_mul_inv, mul_assoc, mul_comm c⁻¹] diff --git a/Mathlib/Algebra/GroupWithZero/Commute.lean b/Mathlib/Algebra/GroupWithZero/Commute.lean index 59423cb09016f..24a2a0afe674b 100644 --- a/Mathlib/Algebra/GroupWithZero/Commute.lean +++ b/Mathlib/Algebra/GroupWithZero/Commute.lean @@ -14,7 +14,7 @@ import Mathlib.Tactic.Nontriviality assert_not_exists DenselyOrdered -variable {α M₀ G₀ M₀' G₀' F F' : Type*} +variable {M₀ G₀ : Type*} variable [MonoidWithZero M₀] namespace Ring @@ -83,7 +83,7 @@ theorem div_left (hac : Commute a c) (hbc : Commute b c) : Commute (a / b) c := end Commute section GroupWithZero -variable {G₀ : Type*} [GroupWithZero G₀] {a : G₀} {m n : ℕ} +variable {G₀ : Type*} [GroupWithZero G₀] theorem pow_inv_comm₀ (a : G₀) (m n : ℕ) : a⁻¹ ^ m * a ^ n = a ^ n * a⁻¹ ^ m := (Commute.refl a).inv_left₀.pow_pow m n diff --git a/Mathlib/Algebra/GroupWithZero/InjSurj.lean b/Mathlib/Algebra/GroupWithZero/InjSurj.lean index d7a7ff367118a..26278384167c1 100644 --- a/Mathlib/Algebra/GroupWithZero/InjSurj.lean +++ b/Mathlib/Algebra/GroupWithZero/InjSurj.lean @@ -19,7 +19,7 @@ variable {M₀ G₀ M₀' G₀' : Type*} section MulZeroClass -variable [MulZeroClass M₀] {a b : M₀} +variable [MulZeroClass M₀] /-- Pull back a `MulZeroClass` instance along an injective function. See note [reducible non-instances]. -/ diff --git a/Mathlib/Algebra/GroupWithZero/Semiconj.lean b/Mathlib/Algebra/GroupWithZero/Semiconj.lean index f3798456188e4..0dfc56212073d 100644 --- a/Mathlib/Algebra/GroupWithZero/Semiconj.lean +++ b/Mathlib/Algebra/GroupWithZero/Semiconj.lean @@ -13,7 +13,7 @@ import Mathlib.Algebra.Group.Semiconj.Units assert_not_exists DenselyOrdered -variable {α M₀ G₀ M₀' G₀' F F' : Type*} +variable {G₀ : Type*} namespace SemiconjBy diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index c413ab0a165b8..bbf5c0a8dc88f 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -94,7 +94,6 @@ def invMonoidWithZeroHom {G₀ : Type*} [CommGroupWithZero G₀] : G₀ →*₀ namespace Units variable [GroupWithZero G₀] -variable {a b : G₀} @[simp] theorem smul_mk0 {α : Type*} [SMul G₀ α] {g : G₀} (hg : g ≠ 0) (a : α) : mk0 g hg • a = g • a := diff --git a/Mathlib/Algebra/Module/BigOperators.lean b/Mathlib/Algebra/Module/BigOperators.lean index 8b68d57357fd4..d15e7274ec57a 100644 --- a/Mathlib/Algebra/Module/BigOperators.lean +++ b/Mathlib/Algebra/Module/BigOperators.lean @@ -15,7 +15,7 @@ variable {ι κ α β R M : Type*} section AddCommMonoid -variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x y : M) +variable [Semiring R] [AddCommMonoid M] [Module R M] theorem List.sum_smul {l : List R} {x : M} : l.sum • x = (l.map fun r ↦ r • x).sum := map_list_sum ((smulAddHom R M).flip x) l diff --git a/Mathlib/Algebra/Polynomial/Induction.lean b/Mathlib/Algebra/Polynomial/Induction.lean index 4ba0e2bf397db..a4fc244b867ab 100644 --- a/Mathlib/Algebra/Polynomial/Induction.lean +++ b/Mathlib/Algebra/Polynomial/Induction.lean @@ -26,12 +26,11 @@ open Polynomial universe u v w x y z -variable {R : Type u} {S : Type v} {T : Type w} {ι : Type x} {k : Type y} {A : Type z} {a b : R} - {m n : ℕ} +variable {R : Type u} section Semiring -variable [Semiring R] {p q r : R[X]} +variable [Semiring R] @[elab_as_elim] protected theorem induction_on {M : R[X] → Prop} (p : R[X]) (h_C : ∀ a, M (C a)) diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index 4ae5170c85619..fd10db100aadd 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -40,13 +40,6 @@ def mulRight [Distrib R] (r : R) : AddHom R R where end AddHom -section AddHomClass - -variable {α β F : Type*} [NonAssocSemiring α] [NonAssocSemiring β] - [FunLike F α β] [AddHomClass F α β] - -end AddHomClass - namespace AddMonoidHom /-- Left multiplication by an element of a (semi)ring is an `AddMonoidHom` -/ @@ -105,7 +98,7 @@ end HasDistribNeg section NonUnitalCommRing -variable {α : Type*} [NonUnitalCommRing α] {a b c : α} +variable {α : Type*} [NonUnitalCommRing α] attribute [local simp] add_assoc add_comm add_left_comm mul_comm diff --git a/Mathlib/Algebra/Ring/Commute.lean b/Mathlib/Algebra/Ring/Commute.lean index 6793e2df1e30d..54cf0da485262 100644 --- a/Mathlib/Algebra/Ring/Commute.lean +++ b/Mathlib/Algebra/Ring/Commute.lean @@ -75,7 +75,7 @@ end section -variable [MulOneClass R] [HasDistribNeg R] {a : R} +variable [MulOneClass R] [HasDistribNeg R] -- Porting note (#10618): no longer needs to be `@[simp]` since `simp` can prove it. -- @[simp] @@ -147,7 +147,7 @@ alias neg_one_pow_two := neg_one_sq end HasDistribNeg section Ring -variable [Ring R] {a b : R} {n : ℕ} +variable [Ring R] {a : R} {n : ℕ} @[simp] lemma neg_one_pow_mul_eq_zero_iff : (-1) ^ n * a = 0 ↔ a = 0 := by rcases neg_one_pow_eq_or R n with h | h <;> simp [h] diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 6ffa90e13d5da..86672a2bb1c98 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -153,7 +153,6 @@ end coe section variable [NonUnitalNonAssocSemiring α] [NonUnitalNonAssocSemiring β] -variable (f : α →ₙ+* β) {x y : α} @[ext] theorem ext ⦃f g : α →ₙ+* β⦄ : (∀ x, f x = g x) → f = g := @@ -225,7 +224,6 @@ theorem coe_comp (g : β →ₙ+* γ) (f : α →ₙ+* β) : ⇑(g.comp f) = g @[simp] theorem comp_apply (g : β →ₙ+* γ) (f : α →ₙ+* β) (x : α) : g.comp f x = g (f x) := rfl -variable (g : β →ₙ+* γ) (f : α →ₙ+* β) @[simp] theorem coe_comp_addMonoidHom (g : β →ₙ+* γ) (f : α →ₙ+* β) : @@ -441,7 +439,7 @@ end coe section -variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} (f : α →+* β) {x y : α} +variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} (f : α →+* β) protected theorem congr_fun {f g : α →+* β} (h : f = g) (x : α) : f x = g x := DFunLike.congr_fun h x diff --git a/Mathlib/Algebra/Ring/Semiconj.lean b/Mathlib/Algebra/Ring/Semiconj.lean index 7ca35cfbf5af9..8b20e68a71a06 100644 --- a/Mathlib/Algebra/Ring/Semiconj.lean +++ b/Mathlib/Algebra/Ring/Semiconj.lean @@ -19,9 +19,9 @@ For the definitions of semirings and rings see `Mathlib.Algebra.Ring.Defs`. -/ -universe u v w x +universe u -variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} +variable {R : Type u} open Function @@ -59,7 +59,7 @@ end section -variable [MulOneClass R] [HasDistribNeg R] {a x y : R} +variable [MulOneClass R] [HasDistribNeg R] -- Porting note: `simpNF` told me to remove `simp` attribute theorem neg_one_right (a : R) : SemiconjBy a (-1) (-1) := diff --git a/Mathlib/Algebra/Ring/Units.lean b/Mathlib/Algebra/Ring/Units.lean index fbcb29f67829f..3e177e160d871 100644 --- a/Mathlib/Algebra/Ring/Units.lean +++ b/Mathlib/Algebra/Ring/Units.lean @@ -49,7 +49,7 @@ end HasDistribNeg section Ring -variable [Ring α] {a b : α} +variable [Ring α] -- Needs to have higher simp priority than divp_add_divp. 1000 is the default priority. @[field_simps 1010] diff --git a/Mathlib/Control/Basic.lean b/Mathlib/Control/Basic.lean index 0df3671f162c4..9f03e34b31c1e 100644 --- a/Mathlib/Control/Basic.lean +++ b/Mathlib/Control/Basic.lean @@ -214,8 +214,6 @@ class CommApplicative (m : Type u → Type v) [Applicative m] extends LawfulAppl open Functor -variable {m} - theorem CommApplicative.commutative_map {m : Type u → Type v} [h : Applicative m] [CommApplicative m] {α β γ} (a : m α) (b : m β) {f : α → β → γ} : f <$> a <*> b = flip f <$> b <*> a := diff --git a/Mathlib/Control/Lawful.lean b/Mathlib/Control/Lawful.lean index 769ec2655f4ec..ce4807b4cde66 100644 --- a/Mathlib/Control/Lawful.lean +++ b/Mathlib/Control/Lawful.lean @@ -48,7 +48,7 @@ end StateT namespace ExceptT -variable {α β ε : Type u} {m : Type u → Type v} (x : ExceptT ε m α) +variable {α ε : Type u} {m : Type u → Type v} (x : ExceptT ε m α) -- Porting note: This is proven by proj reduction in Lean 3. @[simp] diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index 48aaa0d08f165..0583ba40e98cc 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -207,7 +207,7 @@ end Subtype namespace Subtype /-! Some facts about sets, which require that `α` is a type. -/ -variable {α β γ : Type*} {p : α → Prop} +variable {α : Type*} @[simp] theorem coe_prop {S : Set α} (a : { a // a ∈ S }) : ↑a ∈ S := diff --git a/Mathlib/Lean/PrettyPrinter/Delaborator.lean b/Mathlib/Lean/PrettyPrinter/Delaborator.lean index bcc547afa2239..557bdf947198a 100644 --- a/Mathlib/Lean/PrettyPrinter/Delaborator.lean +++ b/Mathlib/Lean/PrettyPrinter/Delaborator.lean @@ -14,13 +14,6 @@ namespace Lean.PrettyPrinter.Delaborator open Lean.Meta Lean.SubExpr SubExpr -namespace SubExpr - -variable {α : Type} [Inhabited α] -variable {m : Type → Type} [Monad m] - -end SubExpr - /-- Assuming the current expression in a lambda or pi, descend into the body using an unused name generated from the binder's name. Provides `d` with both `Syntax` for the bound name as an identifier diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index b33cfc478250f..8fd66baa31e7d 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -435,8 +435,7 @@ end Equality section Quantifiers section Dependent -variable {α : Sort*} {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*} - {ε : ∀ a b c, δ a b c → Sort*} +variable {α : Sort*} {β : α → Sort*} {γ : ∀ a, β a → Sort*} theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) = ∀ a, β' a := (funext h : β = β') ▸ rfl @@ -462,7 +461,7 @@ theorem Exists₃.imp {p q : ∀ a b, γ a b → Prop} (h : ∀ a b c, p a b c end Dependent -variable {α β : Sort*} {p q : α → Prop} +variable {α β : Sort*} {p : α → Prop} theorem forall_swap {p : α → β → Prop} : (∀ x y, p x y) ↔ ∀ y x, p x y := ⟨fun f x y ↦ f y x, fun f x y ↦ f y x⟩ @@ -745,7 +744,7 @@ noncomputable def Exists.classicalRecOn {α : Sort*} {p : α → Prop} (h : ∃ /-! ### Declarations about bounded quantifiers -/ section BoundedQuantifiers -variable {α : Sort*} {r p q : α → Prop} {P Q : ∀ x, p x → Prop} {b : Prop} +variable {α : Sort*} {r p q : α → Prop} {P Q : ∀ x, p x → Prop} theorem bex_def : (∃ (x : _) (_ : p x), q x) ↔ ∃ x, p x ∧ q x := ⟨fun ⟨x, px, qx⟩ ↦ ⟨x, px, qx⟩, fun ⟨x, px, qx⟩ ↦ ⟨x, px, qx⟩⟩ diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d68413363ef4b..d8c4acbd7d587 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1580,8 +1580,6 @@ namespace Equiv section -variable (P : α → Sort w) (e : α ≃ β) - /-- Transport dependent functions through an equivalence of the base space. -/ @[simps apply, simps (config := .lemmasOnly) symm_apply] diff --git a/Mathlib/Logic/Function/Conjugate.lean b/Mathlib/Logic/Function/Conjugate.lean index 76dc5be2b2999..69098eb142cbf 100644 --- a/Mathlib/Logic/Function/Conjugate.lean +++ b/Mathlib/Logic/Function/Conjugate.lean @@ -31,7 +31,7 @@ def Semiconj (f : α → β) (ga : α → α) (gb : β → β) : Prop := namespace Semiconj -variable {f fab : α → β} {fbc : β → γ} {ga ga' : α → α} {gb gb' : β → β} {gc gc' : γ → γ} +variable {f fab : α → β} {fbc : β → γ} {ga ga' : α → α} {gb gb' : β → β} {gc : γ → γ} /-- Definition of `Function.Semiconj` in terms of functional equality. -/ lemma _root_.Function.semiconj_iff_comp_eq : Semiconj f ga gb ↔ f ∘ ga = gb ∘ f := funext_iff.symm diff --git a/Mathlib/Logic/Relator.lean b/Mathlib/Logic/Relator.lean index 7e0f0e054c1ac..83b0b7e2046b8 100644 --- a/Mathlib/Logic/Relator.lean +++ b/Mathlib/Logic/Relator.lean @@ -96,8 +96,8 @@ lemma rel_not : (Iff ⇒ Iff) Not Not := lemma bi_total_eq {α : Type u₁} : Relator.BiTotal (@Eq α) := { left := fun a => ⟨a, rfl⟩, right := fun a => ⟨a, rfl⟩ } -variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} -variable {r : α → β → Prop} {p : β → γ → Prop} {q : γ → δ → Prop} +variable {α : Type*} {β : Type*} {γ : Type*} +variable {r : α → β → Prop} lemma LeftUnique.flip (h : LeftUnique r) : RightUnique (flip r) := fun _ _ _ h₁ h₂ => h h₁ h₂ diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index 145bc1d43f421..678d68e2ff016 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -247,7 +247,6 @@ instance {α} [IsEmpty α] : Unique (Option α) := end Option section Subtype -variable {α : Sort u} instance Unique.subtypeEq (y : α) : Unique { x // x = y } where default := ⟨y, rfl⟩ diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 628c729ea2354..08b21b2de4986 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -186,7 +186,7 @@ end namespace Eq -variable [Preorder α] {x y z : α} +variable [Preorder α] {x y : α} /-- If `x = y` then `y ≤ x`. Note: this lemma uses `y ≤ x` instead of `x ≥ y`, because `le` is used almost exclusively in mathlib. -/ diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index c53c3b994e122..93e79ca56a566 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -552,7 +552,6 @@ lemma max_lt (h₁ : a < c) (h₂ : b < c) : max a b < c := by cases le_total a b <;> simp [max_eq_left, max_eq_right, *] section Ord -variable {o : Ordering} lemma compare_lt_iff_lt : compare a b = .lt ↔ a < b := by rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index 8964e28ae8b2a..53ccb40f79265 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -683,7 +683,7 @@ variable [LinearOrder α] section LocallyFiniteOrder -variable [LocallyFiniteOrder α] {a b : α} +variable [LocallyFiniteOrder α] theorem Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := by @@ -790,7 +790,7 @@ end LinearOrder section Lattice -variable [Lattice α] [LocallyFiniteOrder α] {a a₁ a₂ b b₁ b₂ c x : α} +variable [Lattice α] [LocallyFiniteOrder α] {a a₁ a₂ b b₁ b₂ x : α} theorem uIcc_toDual (a b : α) : [[toDual a, toDual b]] = [[a, b]].map toDual.toEmbedding := Icc_toDual _ _ @@ -862,7 +862,7 @@ end Lattice section DistribLattice -variable [DistribLattice α] [LocallyFiniteOrder α] {a a₁ a₂ b b₁ b₂ c x : α} +variable [DistribLattice α] [LocallyFiniteOrder α] {a b c : α} theorem eq_of_mem_uIcc_of_mem_uIcc : a ∈ [[b, c]] → b ∈ [[a, c]] → a = b := by simp_rw [mem_uIcc] @@ -883,7 +883,7 @@ end DistribLattice section LinearOrder -variable [LinearOrder α] [LocallyFiniteOrder α] {a a₁ a₂ b b₁ b₂ c x : α} +variable [LinearOrder α] [LocallyFiniteOrder α] {a a₁ a₂ b b₁ b₂ c : α} theorem Icc_min_max : Icc (min a b) (max a b) = [[a, b]] := rfl diff --git a/Mathlib/Order/Max.lean b/Mathlib/Order/Max.lean index b355050f7aea9..61f5b1037d505 100644 --- a/Mathlib/Order/Max.lean +++ b/Mathlib/Order/Max.lean @@ -356,7 +356,7 @@ end PartialOrder section Prod -variable [Preorder α] [Preorder β] {a a₁ a₂ : α} {b b₁ b₂ : β} {x y : α × β} +variable [Preorder α] [Preorder β] {a : α} {b : β} {x : α × β} theorem IsBot.prod_mk (ha : IsBot a) (hb : IsBot b) : IsBot (a, b) := fun _ => ⟨ha _, hb _⟩ diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index d8d173cca893a..aef647a6f8d31 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -927,7 +927,7 @@ lemma gc_pred_succ : GaloisConnection (pred : α → α) succ := fun _ _ ↦ pre end Preorder -variable [PartialOrder α] [SuccOrder α] [PredOrder α] {a b : α} +variable [PartialOrder α] [SuccOrder α] [PredOrder α] {a : α} @[simp] theorem succ_pred_of_not_isMin (h : ¬IsMin a) : succ (pred a) = a := diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index f712bc83a22b8..026257adcaf92 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -146,12 +146,12 @@ end ProbabilityTheory namespace MeasureTheory.Measure -variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} +variable {β Ω : Type*} {mβ : MeasurableSpace β} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] section Lintegral -variable [CountableOrCountablyGenerated α β] {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] +variable {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] {f : β × Ω → ℝ≥0∞} lemma lintegral_condKernel_mem {s : Set (β × Ω)} (hs : MeasurableSet s) : diff --git a/Mathlib/Topology/AlexandrovDiscrete.lean b/Mathlib/Topology/AlexandrovDiscrete.lean index b03c9b2cb45eb..517c0d910bfeb 100644 --- a/Mathlib/Topology/AlexandrovDiscrete.lean +++ b/Mathlib/Topology/AlexandrovDiscrete.lean @@ -114,8 +114,6 @@ lemma closure_sUnion (S : Set (Set α)) : closure (⋃₀ S) = ⋃ s ∈ S, clos end AlexandrovDiscrete -variable {s t : Set α} {a x y : α} - lemma Inducing.alexandrovDiscrete [AlexandrovDiscrete α] {f : β → α} (h : Inducing f) : AlexandrovDiscrete β where isOpen_sInter S hS := by diff --git a/Mathlib/Topology/MetricSpace/Infsep.lean b/Mathlib/Topology/MetricSpace/Infsep.lean index 027a2e3139f9c..e451f9e7817b3 100644 --- a/Mathlib/Topology/MetricSpace/Infsep.lean +++ b/Mathlib/Topology/MetricSpace/Infsep.lean @@ -171,7 +171,7 @@ end EDist section PseudoEMetricSpace -variable [PseudoEMetricSpace α] {x y z : α} {s t : Set α} +variable [PseudoEMetricSpace α] {x y z : α} {s : Set α} theorem einfsep_pair (hxy : x ≠ y) : ({x, y} : Set α).einfsep = edist x y := by nth_rw 1 [← min_self (edist x y)] @@ -238,7 +238,7 @@ end PseudoMetricSpace section EMetricSpace -variable [EMetricSpace α] {x y z : α} {s t : Set α} {C : ℝ≥0∞} {sC : Set ℝ≥0∞} +variable [EMetricSpace α] {s : Set α} theorem einfsep_pos_of_finite [Finite s] : 0 < s.einfsep := by cases nonempty_fintype s @@ -312,7 +312,7 @@ end EDist section PseudoEMetricSpace -variable [PseudoEMetricSpace α] {x y : α} {s : Set α} +variable [PseudoEMetricSpace α] {x y : α} theorem infsep_pair_eq_toReal : ({x, y} : Set α).infsep = (edist x y).toReal := by by_cases hxy : x = y diff --git a/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean b/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean index 058f2115e3541..b7b081e3fb55d 100644 --- a/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean +++ b/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean @@ -27,7 +27,7 @@ open Set Metric open Topology variable {α : Type u} {ι : Type v} [MetricSpace α] [ProperSpace α] {c : ι → α} -variable {x : α} {r : ℝ} {s : Set α} +variable {s : Set α} /-- **Shrinking lemma** for coverings by open balls in a proper metric space. A point-finite open cover of a closed subset of a proper metric space by open balls can be shrunk to a new cover by From 856178f8fbcd2932823b5cfe57d1e669bd294c88 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 30 Sep 2024 02:03:22 +0000 Subject: [PATCH 048/174] chore(Topology): move a folder (#17256) Move `Topology/ContinuousFunction` to `Topology/ContinuousMap` so that it agrees with the type name. This PR was generated by the following commands: ```bash git mv Mathlib/Topology/ContinuousFunction Mathlib/Topology/ContinuousMap git grep -l 'Topology.ContinuousF' | xargs sed -e 's/Topology.ContinuousFunction/Topology.ContinuousMap/' -i lake exe mk_all --git ``` --- Counterexamples/Phillips.lean | 2 +- Mathlib.lean | 38 +++++++++---------- .../NonUnital.lean | 4 +- .../ContinuousFunctionalCalculus/Order.lean | 2 +- .../ContinuousFunctionalCalculus/Unique.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 2 +- .../Analysis/CStarAlgebra/GelfandDuality.lean | 8 ++-- Mathlib/Analysis/Complex/Circle.lean | 2 +- Mathlib/Analysis/Fourier/AddCircle.lean | 2 +- .../Analysis/Normed/Group/CocompactMap.lean | 2 +- .../Analysis/Normed/Group/ZeroAtInfty.lean | 2 +- Mathlib/Analysis/Normed/Lp/LpEquiv.lean | 2 +- .../Analysis/SpecialFunctions/Bernstein.lean | 4 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 2 +- .../Manifold/VectorBundle/SmoothSection.lean | 2 +- .../Matrix/HermitianFunctionalCalculus.lean | 2 +- Mathlib/MeasureTheory/Function/AEEqFun.lean | 2 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- .../Integral/RieszMarkovKakutani.lean | 2 +- .../MeasureTheory/Integral/SetIntegral.lean | 2 +- .../Topology/Algebra/ContinuousAffineMap.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 2 +- Mathlib/Topology/Algebra/Module/Basic.lean | 2 +- .../Algebra/Module/CharacterSpace.lean | 2 +- Mathlib/Topology/Algebra/Monoid.lean | 2 +- Mathlib/Topology/Algebra/Star.lean | 2 +- Mathlib/Topology/Category/TopCat/Basic.lean | 2 +- Mathlib/Topology/CompactOpen.lean | 2 +- .../Algebra.lean | 2 +- .../Basic.lean | 0 .../Bounded.lean | 2 +- .../BoundedCompactlySupported.lean | 2 +- .../CocompactMap.lean | 2 +- .../Compact.lean | 2 +- .../CompactlySupported.lean | 4 +- .../ContinuousMapZero.lean | 4 +- .../Ideals.lean | 2 +- .../LocallyConstant.lean | 4 +- .../Ordered.lean | 2 +- .../Polynomial.lean | 2 +- .../Sigma.lean | 0 .../StarOrdered.lean | 4 +- .../StoneWeierstrass.lean | 4 +- .../T0Sierpinski.lean | 2 +- .../Units.lean | 2 +- .../Weierstrass.lean | 0 .../ZeroAtInfty.lean | 4 +- Mathlib/Topology/Hom/Open.lean | 2 +- Mathlib/Topology/Homotopy/Basic.lean | 2 +- .../MetricSpace/GromovHausdorffRealized.lean | 2 +- .../MetricSpace/ThickenedIndicator.lean | 2 +- Mathlib/Topology/Metrizable/Urysohn.lean | 2 +- Mathlib/Topology/Order/Hom/Basic.lean | 2 +- .../Topology/Order/UpperLowerSetTopology.lean | 2 +- Mathlib/Topology/PartitionOfUnity.lean | 2 +- Mathlib/Topology/Sets/Opens.lean | 2 +- .../Topology/Sheaves/PresheafOfFunctions.lean | 2 +- Mathlib/Topology/Specialization.lean | 2 +- Mathlib/Topology/Spectral/Hom.lean | 2 +- Mathlib/Topology/UrysohnsBounded.lean | 4 +- Mathlib/Topology/UrysohnsLemma.lean | 2 +- test/Continuity.lean | 4 +- 63 files changed, 91 insertions(+), 91 deletions(-) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Algebra.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Basic.lean (100%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Bounded.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/BoundedCompactlySupported.lean (98%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/CocompactMap.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Compact.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/CompactlySupported.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/ContinuousMapZero.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Ideals.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/LocallyConstant.lean (94%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Ordered.lean (98%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Polynomial.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Sigma.lean (100%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/StarOrdered.lean (97%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/StoneWeierstrass.lean (99%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/T0Sierpinski.lean (97%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Units.lean (98%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/Weierstrass.lean (100%) rename Mathlib/Topology/{ContinuousFunction => ContinuousMap}/ZeroAtInfty.lean (99%) diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 86aeacfb5f059..cf03d0dc7c250 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.MeasureTheory.Integral.SetIntegral import Mathlib.MeasureTheory.Measure.Lebesgue.Basic -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded /-! # A counterexample on Pettis integrability diff --git a/Mathlib.lean b/Mathlib.lean index 9f3160338f29b..1c70c47509e0a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4575,25 +4575,25 @@ import Mathlib.Topology.Connected.PathConnected import Mathlib.Topology.Connected.Separation import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Constructions -import Mathlib.Topology.ContinuousFunction.Algebra -import Mathlib.Topology.ContinuousFunction.Basic -import Mathlib.Topology.ContinuousFunction.Bounded -import Mathlib.Topology.ContinuousFunction.BoundedCompactlySupported -import Mathlib.Topology.ContinuousFunction.CocompactMap -import Mathlib.Topology.ContinuousFunction.Compact -import Mathlib.Topology.ContinuousFunction.CompactlySupported -import Mathlib.Topology.ContinuousFunction.ContinuousMapZero -import Mathlib.Topology.ContinuousFunction.Ideals -import Mathlib.Topology.ContinuousFunction.LocallyConstant -import Mathlib.Topology.ContinuousFunction.Ordered -import Mathlib.Topology.ContinuousFunction.Polynomial -import Mathlib.Topology.ContinuousFunction.Sigma -import Mathlib.Topology.ContinuousFunction.StarOrdered -import Mathlib.Topology.ContinuousFunction.StoneWeierstrass -import Mathlib.Topology.ContinuousFunction.T0Sierpinski -import Mathlib.Topology.ContinuousFunction.Units -import Mathlib.Topology.ContinuousFunction.Weierstrass -import Mathlib.Topology.ContinuousFunction.ZeroAtInfty +import Mathlib.Topology.ContinuousMap.Algebra +import Mathlib.Topology.ContinuousMap.Basic +import Mathlib.Topology.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.BoundedCompactlySupported +import Mathlib.Topology.ContinuousMap.CocompactMap +import Mathlib.Topology.ContinuousMap.Compact +import Mathlib.Topology.ContinuousMap.CompactlySupported +import Mathlib.Topology.ContinuousMap.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.Ideals +import Mathlib.Topology.ContinuousMap.LocallyConstant +import Mathlib.Topology.ContinuousMap.Ordered +import Mathlib.Topology.ContinuousMap.Polynomial +import Mathlib.Topology.ContinuousMap.Sigma +import Mathlib.Topology.ContinuousMap.StarOrdered +import Mathlib.Topology.ContinuousMap.StoneWeierstrass +import Mathlib.Topology.ContinuousMap.T0Sierpinski +import Mathlib.Topology.ContinuousMap.Units +import Mathlib.Topology.ContinuousMap.Weierstrass +import Mathlib.Topology.ContinuousMap.ZeroAtInfty import Mathlib.Topology.ContinuousOn import Mathlib.Topology.CountableSeparatingOn import Mathlib.Topology.Covering diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 6a52fbf1fd1e4..3ebb272f0dee6 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Algebra.Algebra.Quasispectrum -import Mathlib.Topology.ContinuousFunction.Compact -import Mathlib.Topology.ContinuousFunction.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.Compact +import Mathlib.Topology.ContinuousMap.ContinuousMapZero import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital import Mathlib.Topology.UniformSpace.CompactConvergence diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 0a89ab198431c..8536353d9475b 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -7,7 +7,7 @@ Authors: Frédéric Dupuis import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow -import Mathlib.Topology.ContinuousFunction.StarOrdered +import Mathlib.Topology.ContinuousMap.StarOrdered /-! # Facts about star-ordered rings that depend on the continuous functional calculus diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 2895a56f474d3..155a1b6dd519c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital -import Mathlib.Topology.ContinuousFunction.StoneWeierstrass +import Mathlib.Topology.ContinuousMap.StoneWeierstrass /-! # Uniqueness of the continuous functional calculus diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index e64b9cdf7d629..89017072beb7c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Algebra.Quasispectrum import Mathlib.Algebra.Algebra.Spectrum import Mathlib.Algebra.Order.Star.Basic import Mathlib.Topology.Algebra.Polynomial -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Tactic.ContinuousFunctionalCalculus /-! diff --git a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean index c3e2791a52a09..af17ea013e3fc 100644 --- a/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean +++ b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean @@ -6,11 +6,11 @@ Authors: Jireh Loreaux import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.Normed.Group.Quotient import Mathlib.Analysis.Normed.Algebra.Basic -import Mathlib.Topology.ContinuousFunction.Units -import Mathlib.Topology.ContinuousFunction.Compact +import Mathlib.Topology.ContinuousMap.Units +import Mathlib.Topology.ContinuousMap.Compact import Mathlib.Topology.Algebra.Algebra -import Mathlib.Topology.ContinuousFunction.Ideals -import Mathlib.Topology.ContinuousFunction.StoneWeierstrass +import Mathlib.Topology.ContinuousMap.Ideals +import Mathlib.Topology.ContinuousMap.StoneWeierstrass /-! # Gelfand Duality diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index bda86fd651ad1..72e66ae46c1b7 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ import Mathlib.Analysis.SpecialFunctions.Exp -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Analysis.Normed.Field.UnitBall /-! diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 86ac394b06e8d..7ca64c6a78543 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -10,7 +10,7 @@ import Mathlib.MeasureTheory.Function.ContinuousMapDense import Mathlib.MeasureTheory.Function.L2Space import Mathlib.MeasureTheory.Group.Integral import Mathlib.MeasureTheory.Integral.Periodic -import Mathlib.Topology.ContinuousFunction.StoneWeierstrass +import Mathlib.Topology.ContinuousMap.StoneWeierstrass import Mathlib.MeasureTheory.Integral.FundThmCalculus /-! diff --git a/Mathlib/Analysis/Normed/Group/CocompactMap.lean b/Mathlib/Analysis/Normed/Group/CocompactMap.lean index ff947f47393ef..40889eae41c45 100644 --- a/Mathlib/Analysis/Normed/Group/CocompactMap.lean +++ b/Mathlib/Analysis/Normed/Group/CocompactMap.lean @@ -5,7 +5,7 @@ Authors: Moritz Doll -/ import Mathlib.Analysis.Normed.Group.Basic -import Mathlib.Topology.ContinuousFunction.CocompactMap +import Mathlib.Topology.ContinuousMap.CocompactMap import Mathlib.Topology.MetricSpace.Bounded /-! diff --git a/Mathlib/Analysis/Normed/Group/ZeroAtInfty.lean b/Mathlib/Analysis/Normed/Group/ZeroAtInfty.lean index a2cabd37f31d1..825bb7c7941a6 100644 --- a/Mathlib/Analysis/Normed/Group/ZeroAtInfty.lean +++ b/Mathlib/Analysis/Normed/Group/ZeroAtInfty.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ -import Mathlib.Topology.ContinuousFunction.ZeroAtInfty +import Mathlib.Topology.ContinuousMap.ZeroAtInfty /-! # ZeroAtInftyContinuousMapClass in normed additive groups diff --git a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean index 8be09ea3634f1..144ad83e536db 100644 --- a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean +++ b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.Normed.Lp.lpSpace import Mathlib.Analysis.Normed.Lp.PiLp -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded /-! # Equivalences among $L^p$ spaces diff --git a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean index 71515ffc2fc55..50a7c0ab65155 100644 --- a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean +++ b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean @@ -5,8 +5,8 @@ Authors: Kim Morrison -/ import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.RingTheory.Polynomial.Bernstein -import Mathlib.Topology.ContinuousFunction.Polynomial -import Mathlib.Topology.ContinuousFunction.Compact +import Mathlib.Topology.ContinuousMap.Polynomial +import Mathlib.Topology.ContinuousMap.Compact /-! # Bernstein approximations and Weierstrass' theorem diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 2986597856fcc..8c7bd3f605070 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -7,7 +7,7 @@ Authors: Frédéric Dupuis import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital -import Mathlib.Topology.ContinuousFunction.StarOrdered +import Mathlib.Topology.ContinuousMap.StarOrdered /-! # The exponential and logarithm based on the continuous functional calculus diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index ff5840676376f..b8beaedc2623e 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, Floris van Doorn -/ import Mathlib.Geometry.Manifold.MFDeriv.Basic -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Geometry.Manifold.Algebra.LieGroup /-! diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 336209008ab7e..ff884d2e289ad 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -7,7 +7,7 @@ Authors: Jon Bannon, Jireh Loreaux import Mathlib.LinearAlgebra.Matrix.Spectrum import Mathlib.LinearAlgebra.Eigenspace.Matrix import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique -import Mathlib.Topology.ContinuousFunction.Units +import Mathlib.Topology.ContinuousMap.Units /-! # Continuous Functional Calculus for Hermitian Matrices diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index 77d4374b0ee83..000e0416082a7 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Zhouhang Zhou -/ import Mathlib.MeasureTheory.Integral.Lebesgue import Mathlib.Order.Filter.Germ.Basic -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic /-! diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index c44471de29253..ed9015bf0ac57 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -13,7 +13,7 @@ import Mathlib.MeasureTheory.Function.LpSeminorm.TriangleInequality import Mathlib.MeasureTheory.Measure.OpenPos import Mathlib.MeasureTheory.Measure.Typeclasses import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace -import Mathlib.Topology.ContinuousFunction.Compact +import Mathlib.Topology.ContinuousMap.Compact import Mathlib.Order.Filter.IndicatorFunction /-! diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 66eafadc4d30a..3e36cb114970b 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Group.Action -import Mathlib.Topology.ContinuousFunction.CocompactMap +import Mathlib.Topology.ContinuousMap.CocompactMap /-! # Measures on Groups diff --git a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean index 20a484f42774c..b6a9c18db7857 100644 --- a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean +++ b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jesse Reimann. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jesse Reimann, Kalle Kytölä -/ -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.Sets.Compacts /-! diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 38f00f6fd9293..4af2e6de396cb 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -7,7 +7,7 @@ import Mathlib.MeasureTheory.Integral.IntegrableOn import Mathlib.MeasureTheory.Integral.Bochner import Mathlib.MeasureTheory.Function.LocallyIntegrable import Mathlib.Topology.MetricSpace.ThickenedIndicator -import Mathlib.Topology.ContinuousFunction.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.ContinuousMapZero import Mathlib.Analysis.NormedSpace.HahnBanach.SeparatingDual /-! diff --git a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean index 39cb34e5d4264..b2d7477d2d987 100644 --- a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean +++ b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.LinearAlgebra.AffineSpace.AffineMap -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Algebra.Module.Basic /-! diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index de9f0789c24cb..00b7d7e857345 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -5,7 +5,7 @@ Authors: Thomas Browning -/ import Mathlib.Topology.Algebra.Equicontinuity import Mathlib.Topology.Algebra.Group.Compact -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.UniformSpace.Ascoli /-! diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index ce5bbbc74d460..36ed9e6021dea 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -7,7 +7,7 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Fréd import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Algebra.UniformGroup -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Algebra.Algebra.Defs import Mathlib.LinearAlgebra.Projection diff --git a/Mathlib/Topology/Algebra/Module/CharacterSpace.lean b/Mathlib/Topology/Algebra/Module/CharacterSpace.lean index aa4aeffcdfcb1..39d280cac3ba7 100644 --- a/Mathlib/Topology/Algebra/Module/CharacterSpace.lean +++ b/Mathlib/Topology/Algebra/Module/CharacterSpace.lean @@ -5,7 +5,7 @@ Authors: Frédéric Dupuis -/ import Mathlib.Topology.Algebra.Module.WeakDual import Mathlib.Algebra.Algebra.Spectrum -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Data.Set.Lattice /-! diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index c2bebad5ed312..a8001b461554b 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.Pointwise import Mathlib.Topology.Algebra.MulAction import Mathlib.Algebra.BigOperators.Pi -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Algebra.Group.ULift /-! diff --git a/Mathlib/Topology/Algebra/Star.lean b/Mathlib/Topology/Algebra/Star.lean index deb57a790041b..abbca7eae5b32 100644 --- a/Mathlib/Topology/Algebra/Star.lean +++ b/Mathlib/Topology/Algebra/Star.lean @@ -6,7 +6,7 @@ Authors: Eric Wieser import Mathlib.Algebra.Star.Pi import Mathlib.Algebra.Star.Prod import Mathlib.Topology.Algebra.Constructions -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Continuity of `star` diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 40459bf75dc71..9addb98e719ef 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Kim Morrison, Mario Carneiro -/ import Mathlib.CategoryTheory.ConcreteCategory.BundledHom -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Category instance for topological spaces diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 99e06e4f3aa34..6b8b401f5542a 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -3,7 +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.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # The compact-open topology diff --git a/Mathlib/Topology/ContinuousFunction/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/Algebra.lean rename to Mathlib/Topology/ContinuousMap/Algebra.lean index f355d40dd16c6..a31ec711b86dc 100644 --- a/Mathlib/Topology/ContinuousFunction/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -14,7 +14,7 @@ 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.ContinuousFunction.Ordered +import Mathlib.Topology.ContinuousMap.Ordered import Mathlib.Topology.UniformSpace.CompactConvergence /-! diff --git a/Mathlib/Topology/ContinuousFunction/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean similarity index 100% rename from Mathlib/Topology/ContinuousFunction/Basic.lean rename to Mathlib/Topology/ContinuousMap/Basic.lean diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousMap/Bounded.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/Bounded.lean rename to Mathlib/Topology/ContinuousMap/Bounded.lean index 9b505b1746655..205125c91bfb4 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Mario Carneiro, Yury Kudryashov, Heather Macbeth -/ import Mathlib.Algebra.Module.MinimalAxioms -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic import Mathlib.Analysis.CStarAlgebra.Basic diff --git a/Mathlib/Topology/ContinuousFunction/BoundedCompactlySupported.lean b/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean similarity index 98% rename from Mathlib/Topology/ContinuousFunction/BoundedCompactlySupported.lean rename to Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean index 9b6c91a69de4e..ff7efe16ef944 100644 --- a/Mathlib/Topology/ContinuousFunction/BoundedCompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yoh Tanimoto. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yoh Tanimoto -/ -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.RingTheory.TwoSidedIdeal.Lattice /-! diff --git a/Mathlib/Topology/ContinuousFunction/CocompactMap.lean b/Mathlib/Topology/ContinuousMap/CocompactMap.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/CocompactMap.lean rename to Mathlib/Topology/ContinuousMap/CocompactMap.lean index 8794fe4397f20..d5662c0299843 100644 --- a/Mathlib/Topology/ContinuousFunction/CocompactMap.lean +++ b/Mathlib/Topology/ContinuousMap/CocompactMap.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.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Cocompact continuous maps diff --git a/Mathlib/Topology/ContinuousFunction/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/Compact.lean rename to Mathlib/Topology/ContinuousMap/Compact.lean index 50b42bb56fb90..eb72857ce4102 100644 --- a/Mathlib/Topology/ContinuousFunction/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.UniformSpace.Compact import Mathlib.Topology.CompactOpen import Mathlib.Topology.Sets.Compacts diff --git a/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/CompactlySupported.lean rename to Mathlib/Topology/ContinuousMap/CompactlySupported.lean index d8d72123235af..2b51d3ea2f2d0 100644 --- a/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Yoh Tanimoto. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yoh Tanimoto -/ -import Mathlib.Topology.ContinuousFunction.CocompactMap -import Mathlib.Topology.ContinuousFunction.ZeroAtInfty +import Mathlib.Topology.ContinuousMap.CocompactMap +import Mathlib.Topology.ContinuousMap.ZeroAtInfty import Mathlib.Topology.Support /-! diff --git a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean rename to Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index a0e26115b77b1..fd74c5bccfa7c 100644 --- a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -3,8 +3,8 @@ 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.Topology.ContinuousFunction.Algebra -import Mathlib.Topology.ContinuousFunction.Compact +import Mathlib.Topology.ContinuousMap.Algebra +import Mathlib.Topology.ContinuousMap.Compact /-! # Continuous maps sending zero to zero diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousMap/Ideals.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/Ideals.lean rename to Mathlib/Topology/ContinuousMap/Ideals.lean index e40d94ca79034..ead82fefc127c 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousMap/Ideals.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Topology.Algebra.Algebra -import Mathlib.Topology.ContinuousFunction.Compact +import Mathlib.Topology.ContinuousMap.Compact import Mathlib.Topology.UrysohnsLemma import Mathlib.Analysis.RCLike.Basic import Mathlib.Analysis.Normed.Ring.Units diff --git a/Mathlib/Topology/ContinuousFunction/LocallyConstant.lean b/Mathlib/Topology/ContinuousMap/LocallyConstant.lean similarity index 94% rename from Mathlib/Topology/ContinuousFunction/LocallyConstant.lean rename to Mathlib/Topology/ContinuousMap/LocallyConstant.lean index a9858942c825d..dce7a9a92d359 100644 --- a/Mathlib/Topology/ContinuousFunction/LocallyConstant.lean +++ b/Mathlib/Topology/ContinuousMap/LocallyConstant.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Topology.LocallyConstant.Algebra -import Mathlib.Topology.ContinuousFunction.Basic -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Basic +import Mathlib.Topology.ContinuousMap.Algebra /-! # The algebra morphism from locally constant functions to continuous functions. diff --git a/Mathlib/Topology/ContinuousFunction/Ordered.lean b/Mathlib/Topology/ContinuousMap/Ordered.lean similarity index 98% rename from Mathlib/Topology/ContinuousFunction/Ordered.lean rename to Mathlib/Topology/ContinuousMap/Ordered.lean index 383451ecd6565..24c45724b82b7 100644 --- a/Mathlib/Topology/ContinuousFunction/Ordered.lean +++ b/Mathlib/Topology/ContinuousMap/Ordered.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Shing Tak Lam -/ -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Order.Lattice import Mathlib.Topology.Order.ProjIcc diff --git a/Mathlib/Topology/ContinuousFunction/Polynomial.lean b/Mathlib/Topology/ContinuousMap/Polynomial.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/Polynomial.lean rename to Mathlib/Topology/ContinuousMap/Polynomial.lean index 5e3cd45660fa2..83a3dd5e6a0ed 100644 --- a/Mathlib/Topology/ContinuousFunction/Polynomial.lean +++ b/Mathlib/Topology/ContinuousMap/Polynomial.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Topology.Algebra.Polynomial -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.UnitInterval import Mathlib.Algebra.Star.Subalgebra diff --git a/Mathlib/Topology/ContinuousFunction/Sigma.lean b/Mathlib/Topology/ContinuousMap/Sigma.lean similarity index 100% rename from Mathlib/Topology/ContinuousFunction/Sigma.lean rename to Mathlib/Topology/ContinuousMap/Sigma.lean diff --git a/Mathlib/Topology/ContinuousFunction/StarOrdered.lean b/Mathlib/Topology/ContinuousMap/StarOrdered.lean similarity index 97% rename from Mathlib/Topology/ContinuousFunction/StarOrdered.lean rename to Mathlib/Topology/ContinuousMap/StarOrdered.lean index 18882fe259d74..b894be96cbeb1 100644 --- a/Mathlib/Topology/ContinuousFunction/StarOrdered.lean +++ b/Mathlib/Topology/ContinuousMap/StarOrdered.lean @@ -5,8 +5,8 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.Complex.Basic import Mathlib.Data.Real.StarOrdered -import Mathlib.Topology.ContinuousFunction.Algebra -import Mathlib.Topology.ContinuousFunction.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.Algebra +import Mathlib.Topology.ContinuousMap.ContinuousMapZero /-! # Continuous functions as a star-ordered ring -/ diff --git a/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean rename to Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index 75dd242760f15..ff47e32977e60 100644 --- a/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -6,8 +6,8 @@ Authors: Kim Morrison, Heather Macbeth import Mathlib.Algebra.Algebra.Subalgebra.Unitization import Mathlib.Analysis.RCLike.Basic import Mathlib.Topology.Algebra.StarSubalgebra -import Mathlib.Topology.ContinuousFunction.ContinuousMapZero -import Mathlib.Topology.ContinuousFunction.Weierstrass +import Mathlib.Topology.ContinuousMap.ContinuousMapZero +import Mathlib.Topology.ContinuousMap.Weierstrass /-! # The Stone-Weierstrass theorem diff --git a/Mathlib/Topology/ContinuousFunction/T0Sierpinski.lean b/Mathlib/Topology/ContinuousMap/T0Sierpinski.lean similarity index 97% rename from Mathlib/Topology/ContinuousFunction/T0Sierpinski.lean rename to Mathlib/Topology/ContinuousMap/T0Sierpinski.lean index 48655f0023ded..4e9903ccc7400 100644 --- a/Mathlib/Topology/ContinuousFunction/T0Sierpinski.lean +++ b/Mathlib/Topology/ContinuousMap/T0Sierpinski.lean @@ -5,7 +5,7 @@ Authors: Ivan Sadofschi Costa -/ import Mathlib.Topology.Order import Mathlib.Topology.Sets.Opens -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Any T0 space embeds in a product of copies of the Sierpinski space. diff --git a/Mathlib/Topology/ContinuousFunction/Units.lean b/Mathlib/Topology/ContinuousMap/Units.lean similarity index 98% rename from Mathlib/Topology/ContinuousFunction/Units.lean rename to Mathlib/Topology/ContinuousMap/Units.lean index a8c880cf3a2bc..0337bcc434faa 100644 --- a/Mathlib/Topology/ContinuousFunction/Units.lean +++ b/Mathlib/Topology/ContinuousMap/Units.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Algebra.Algebra.Spectrum -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra /-! # Units of continuous functions diff --git a/Mathlib/Topology/ContinuousFunction/Weierstrass.lean b/Mathlib/Topology/ContinuousMap/Weierstrass.lean similarity index 100% rename from Mathlib/Topology/ContinuousFunction/Weierstrass.lean rename to Mathlib/Topology/ContinuousMap/Weierstrass.lean diff --git a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean rename to Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean index 932426106e17d..c8c2dd96d5d02 100644 --- a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean @@ -3,8 +3,8 @@ 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.Topology.ContinuousFunction.Bounded -import Mathlib.Topology.ContinuousFunction.CocompactMap +import Mathlib.Topology.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.CocompactMap /-! # Continuous functions vanishing at infinity diff --git a/Mathlib/Topology/Hom/Open.lean b/Mathlib/Topology/Hom/Open.lean index bec44e93eff1e..86051435726a8 100644 --- a/Mathlib/Topology/Hom/Open.lean +++ b/Mathlib/Topology/Hom/Open.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.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Continuous open maps diff --git a/Mathlib/Topology/Homotopy/Basic.lean b/Mathlib/Topology/Homotopy/Basic.lean index caf0e3a6983a5..e736983bd8031 100644 --- a/Mathlib/Topology/Homotopy/Basic.lean +++ b/Mathlib/Topology/Homotopy/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Shing Tak Lam -/ import Mathlib.Topology.Order.ProjIcc -import Mathlib.Topology.ContinuousFunction.Ordered +import Mathlib.Topology.ContinuousMap.Ordered import Mathlib.Topology.CompactOpen import Mathlib.Topology.UnitInterval diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index f343b3d1c3901..9a6979a644bb1 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Topology.MetricSpace.Gluing import Mathlib.Topology.MetricSpace.HausdorffDistance -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded /-! # The Gromov-Hausdorff distance is realized diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index c274fbbc8e376..37a5faee8383a 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ import Mathlib.Data.ENNReal.Basic -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.MetricSpace.Thickening /-! diff --git a/Mathlib/Topology/Metrizable/Urysohn.lean b/Mathlib/Topology/Metrizable/Urysohn.lean index f136cbf094264..bddb6a6ead668 100644 --- a/Mathlib/Topology/Metrizable/Urysohn.lean +++ b/Mathlib/Topology/Metrizable/Urysohn.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Topology.UrysohnsLemma -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.Metrizable.Basic /-! # Urysohn's Metrization Theorem diff --git a/Mathlib/Topology/Order/Hom/Basic.lean b/Mathlib/Topology/Order/Hom/Basic.lean index 11dc84c396d3c..94dd1b5acb371 100644 --- a/Mathlib/Topology/Order/Hom/Basic.lean +++ b/Mathlib/Topology/Order/Hom/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.Hom.Basic -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Continuous order homomorphisms diff --git a/Mathlib/Topology/Order/UpperLowerSetTopology.lean b/Mathlib/Topology/Order/UpperLowerSetTopology.lean index 5aa350f222cbb..02681dbd9f86d 100644 --- a/Mathlib/Topology/Order/UpperLowerSetTopology.lean +++ b/Mathlib/Topology/Order/UpperLowerSetTopology.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ import Mathlib.Topology.AlexandrovDiscrete -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Order.LowerUpperTopology /-! diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index bdc558694b1cd..e44e945ac68f0 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.BigOperators.Finprod -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.Compactness.Paracompact import Mathlib.Topology.ShrinkingLemma import Mathlib.Topology.UrysohnsLemma diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 5e2b0453cc80e..cf7ecc4c6353c 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn import Mathlib.Order.Hom.CompleteLattice import Mathlib.Topology.Bases import Mathlib.Topology.Homeomorph -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Order.CompactlyGenerated.Basic import Mathlib.Order.Copy diff --git a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean index 23061bd8b782f..6049ed866f9d8 100644 --- a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean +++ b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean @@ -6,7 +6,7 @@ Authors: Kim Morrison import Mathlib.CategoryTheory.Yoneda import Mathlib.Topology.Sheaves.Presheaf import Mathlib.Topology.Category.TopCommRingCat -import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousMap.Algebra /-! # Presheaves of functions diff --git a/Mathlib/Topology/Specialization.lean b/Mathlib/Topology/Specialization.lean index 874083b9a9983..0cec4f5ed17b5 100644 --- a/Mathlib/Topology/Specialization.lean +++ b/Mathlib/Topology/Specialization.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Order.Category.Preord import Mathlib.Topology.Category.TopCat.Basic -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.Separation import Mathlib.Topology.Order.UpperLowerSetTopology diff --git a/Mathlib/Topology/Spectral/Hom.lean b/Mathlib/Topology/Spectral/Hom.lean index 078b9875e0dbe..30c704a2d0504 100644 --- a/Mathlib/Topology/Spectral/Hom.lean +++ b/Mathlib/Topology/Spectral/Hom.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.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic /-! # Spectral maps diff --git a/Mathlib/Topology/UrysohnsBounded.lean b/Mathlib/Topology/UrysohnsBounded.lean index fa436cb53dfbd..16835df30f181 100644 --- a/Mathlib/Topology/UrysohnsBounded.lean +++ b/Mathlib/Topology/UrysohnsBounded.lean @@ -4,14 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Topology.UrysohnsLemma -import Mathlib.Topology.ContinuousFunction.Bounded +import Mathlib.Topology.ContinuousMap.Bounded /-! # Urysohn's lemma for bounded continuous functions In this file we reformulate Urysohn's lemma `exists_continuous_zero_one_of_isClosed` in terms of bounded continuous functions `X →ᵇ ℝ`. These lemmas live in a separate file because -`Topology.ContinuousFunction.Bounded` imports too many other files. +`Topology.ContinuousMap.Bounded` imports too many other files. ## Tags diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 4c07760f23a50..705eb5ca08358 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.Ordered -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Topology.GDelta import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.SpecificLimits.Basic diff --git a/test/Continuity.lean b/test/Continuity.lean index 7df66051c1b4e..70da467d763ee 100644 --- a/test/Continuity.lean +++ b/test/Continuity.lean @@ -1,6 +1,6 @@ import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic import Mathlib.Topology.Basic -import Mathlib.Topology.ContinuousFunction.Basic +import Mathlib.Topology.ContinuousMap.Basic set_option autoImplicit true section basic @@ -45,7 +45,7 @@ example : Continuous (fun x : ℝ => exp ((max x (-x)) + sin x)^2) := by example : Continuous (fun x : ℝ => exp ((max x (-x)) + sin (cos x))^2) := by continuity --- Examples taken from `Topology.ContinuousFunction.Basic`: +-- Examples taken from `Topology.ContinuousMap.Basic`: example (b : Y) : Continuous (fun _ : X => b) := by continuity From 9a1a45ad35bef46fe855d469b7283397357b10a9 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Mon, 30 Sep 2024 02:32:54 +0000 Subject: [PATCH 049/174] feat (Topology/Instances/EReal): `EReal.toReal` tends to top at top (#17100) - Add `nhdsWithin_top` and `nhdsWithin_bot`: the punctured neighbourhoods of top and bot in `EReal` are the map of `atTop` and `atBot`. - Add `tendsto_toReal_atTop` and `tendsto_toReal_atBot`. --- Mathlib/Topology/Instances/EReal.lean | 35 +++++++++++++++++++++++++++ 1 file changed, 35 insertions(+) diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 5ad8821215f37..8b4ae160970e4 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -145,6 +145,41 @@ theorem tendsto_nhds_bot_iff_real {α : Type*} {m : α → EReal} {f : Filter α Tendsto m f (𝓝 ⊥) ↔ ∀ x : ℝ, ∀ᶠ a in f, m a < x := nhds_bot_basis.tendsto_right_iff.trans <| by simp only [true_implies, mem_Iio] +lemma nhdsWithin_top : 𝓝[≠] (⊤ : EReal) = (atTop).map Real.toEReal := by + apply (nhdsWithin_hasBasis nhds_top_basis_Ici _).ext (atTop_basis.map Real.toEReal) + · simp only [EReal.image_coe_Ici, true_and] + intro x hx + by_cases hx_bot : x = ⊥ + · simp [hx_bot] + lift x to ℝ using ⟨hx.ne_top, hx_bot⟩ + refine ⟨x, fun x ⟨h1, h2⟩ ↦ ?_⟩ + simp [h1, h2.ne_top] + · simp only [EReal.image_coe_Ici, true_implies] + refine fun x ↦ ⟨x, ⟨EReal.coe_lt_top x, fun x ⟨(h1 : _ ≤ x), h2⟩ ↦ ?_⟩⟩ + simp [h1, Ne.lt_top' fun a ↦ h2 a.symm] + +lemma nhdsWithin_bot : 𝓝[≠] (⊥ : EReal) = (atBot).map Real.toEReal := by + apply (nhdsWithin_hasBasis nhds_bot_basis_Iic _).ext (atBot_basis.map Real.toEReal) + · simp only [EReal.image_coe_Iic, Set.subset_compl_singleton_iff, Set.mem_Ioc, lt_self_iff_false, + bot_le, and_true, not_false_eq_true, true_and] + intro x hx + by_cases hx_top : x = ⊤ + · simp [hx_top] + lift x to ℝ using ⟨hx_top, hx.ne_bot⟩ + refine ⟨x, fun x ⟨h1, h2⟩ ↦ ?_⟩ + simp [h2, h1.ne_bot] + · simp only [EReal.image_coe_Iic, true_implies] + refine fun x ↦ ⟨x, ⟨EReal.bot_lt_coe x, fun x ⟨(h1 : x ≤ _), h2⟩ ↦ ?_⟩⟩ + simp [h1, Ne.bot_lt' fun a ↦ h2 a.symm] + +lemma tendsto_toReal_atTop : Tendsto EReal.toReal (𝓝[≠] ⊤) atTop := by + rw [nhdsWithin_top, tendsto_map'_iff] + exact tendsto_id + +lemma tendsto_toReal_atBot : Tendsto EReal.toReal (𝓝[≠] ⊥) atBot := by + rw [nhdsWithin_bot, tendsto_map'_iff] + exact tendsto_id + /-! ### Infs and Sups -/ variable {α : Type*} {u v : α → EReal} From ab3370981d777b54aaa244c64143abcb00e9d7a1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 30 Sep 2024 02:32:55 +0000 Subject: [PATCH 050/174] chore(Associated): use `M`, `N` for type variables (#17115) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replace `α` → `M`, `β` → `N` for type variables. --- Mathlib/Algebra/Associated/Basic.lean | 438 +++++++++++++------------- 1 file changed, 219 insertions(+), 219 deletions(-) diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 7eaa44c44512d..e03847e690fc0 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -31,20 +31,20 @@ and prove basic properties of this quotient. assert_not_exists OrderedCommMonoid assert_not_exists Multiset -variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} +variable {M N : Type*} section Prime -variable [CommMonoidWithZero α] +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 : α) : Prop := +def Prime (p : M) : Prop := p ≠ 0 ∧ ¬IsUnit p ∧ ∀ a b, p ∣ a * b → p ∣ a ∨ p ∣ b namespace Prime -variable {p : α} (hp : Prime p) +variable {p : M} (hp : Prime p) include hp theorem ne_zero : p ≠ 0 := @@ -58,19 +58,19 @@ theorem not_dvd_one : ¬p ∣ 1 := theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) -theorem dvd_or_dvd {a b : α} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := +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 : α} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := +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 : α} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := +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 : α} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by +theorem dvd_of_dvd_pow {a : M} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by induction n with | zero => rw [pow_zero] at h @@ -83,24 +83,24 @@ theorem dvd_of_dvd_pow {a : α} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by · assumption · exact ih dvd_pow -theorem dvd_pow_iff_dvd {a : α} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := +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 : α) := fun h => h.ne_zero rfl +theorem not_prime_zero : ¬Prime (0 : M) := fun h => h.ne_zero rfl @[simp] -theorem not_prime_one : ¬Prime (1 : α) := fun h => h.not_unit isUnit_one +theorem not_prime_one : ¬Prime (1 : M) := fun h => h.not_unit isUnit_one section Map -variable [CommMonoidWithZero β] {F : Type*} {G : Type*} [FunLike F α β] -variable [MonoidWithZeroHomClass F α β] [FunLike G β α] [MulHomClass G β α] -variable (f : F) (g : G) {p : α} +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 : β) = a) (hp : Prime (f p)) : Prime p := +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 @@ -110,7 +110,7 @@ theorem comap_prime (hinv : ∀ a, g (f a : β) = a) (hp : Prime (f p)) : Prime · intro h convert ← map_dvd g h <;> apply hinv⟩ -theorem MulEquiv.prime_iff (e : α ≃* β) : Prime p ↔ Prime (e p) := +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⟩ @@ -118,15 +118,15 @@ end Map end Prime -theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero α] {p : α} (hp : Prime p) - {a b : α} : a ∣ p * b → p ∣ a ∨ a ∣ b := by +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 α] {p a b : α} (hp : Prime p) +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 => @@ -138,12 +138,12 @@ theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero α] {p a b : α} 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 α] {p a b : α} (hp : Prime p) +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 α] {p a b : α} +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 @@ -161,7 +161,7 @@ theorem Prime.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd [CancelCommMonoidWith rw [pow_two, ← mul_assoc] exact dvd_mul_right _ _ -theorem prime_pow_succ_dvd_mul {α : Type*} [CancelCommMonoidWithZero α] {p x y : α} (h : Prime p) +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 @@ -178,7 +178,7 @@ theorem prime_pow_succ_dvd_mul {α : Type*} [CancelCommMonoidWithZero α] {p x y 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 α] (p : α) : Prop where +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 -/ @@ -186,38 +186,38 @@ structure Irreducible [Monoid α] (p : α) : Prop where namespace Irreducible -theorem not_dvd_one [CommMonoid α] {p : α} (hp : Irreducible p) : ¬p ∣ 1 := +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 α] {p : α} (hp : Irreducible p) {a b : α} (h : p = a * b) : +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 α] {p : α} : +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 α] : ¬Irreducible (1 : α) := by simp [irreducible_iff] +theorem not_irreducible_one [Monoid M] : ¬Irreducible (1 : M) := by simp [irreducible_iff] -theorem Irreducible.ne_one [Monoid α] : ∀ {p : α}, Irreducible p → p ≠ 1 +theorem Irreducible.ne_one [Monoid M] : ∀ {p : M}, Irreducible p → p ≠ 1 | _, hp, rfl => not_irreducible_one hp @[simp] -theorem not_irreducible_zero [MonoidWithZero α] : ¬Irreducible (0 : α) +theorem not_irreducible_zero [MonoidWithZero M] : ¬Irreducible (0 : M) | ⟨hn0, h⟩ => - have : IsUnit (0 : α) ∨ IsUnit (0 : α) := h 0 0 (mul_zero 0).symm + have : IsUnit (0 : M) ∨ IsUnit (0 : M) := h 0 0 (mul_zero 0).symm this.elim hn0 hn0 -theorem Irreducible.ne_zero [MonoidWithZero α] : ∀ {p : α}, Irreducible p → p ≠ 0 +theorem Irreducible.ne_zero [MonoidWithZero M] : ∀ {p : M}, Irreducible p → p ≠ 0 | _, hp, rfl => not_irreducible_zero hp -theorem of_irreducible_mul {α} [Monoid α] {x y : α} : Irreducible (x * y) → IsUnit x ∨ IsUnit y +theorem of_irreducible_mul {M} [Monoid M] {x y : M} : Irreducible (x * y) → IsUnit x ∨ IsUnit y | ⟨_, h⟩ => h _ _ rfl -theorem not_irreducible_pow {α} [Monoid α] {x : α} {n : ℕ} (hn : n ≠ 1) : +theorem not_irreducible_pow {M} [Monoid M] {x : M} {n : ℕ} (hn : n ≠ 1) : ¬ Irreducible (x ^ n) := by cases n with | zero => simp @@ -227,7 +227,7 @@ theorem not_irreducible_pow {α} [Monoid α] {x : α} {n : ℕ} (hn : n ≠ 1) : rw [isUnit_pow_iff (Nat.succ_ne_succ.mp hn), or_self] at this exact h₁ (this.pow _) -theorem irreducible_or_factor {α} [Monoid α] (x : α) (h : ¬IsUnit x) : +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 => ?_ @@ -239,20 +239,20 @@ theorem irreducible_or_factor {α} [Monoid α] (x : α) (h : ¬IsUnit x) : 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 α] {p q : α} (hp : Irreducible p) (hq : Irreducible q) : +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 α] {p q : α} (hp : Irreducible p) (hq : Irreducible q) : +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 -variable [Monoid α] +variable [Monoid M] -theorem irreducible_units_mul (a : αˣ) (b : α) : Irreducible (↑a * b) ↔ Irreducible b := by +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] @@ -262,11 +262,11 @@ theorem irreducible_units_mul (a : αˣ) (b : α) : Irreducible (↑a * b) ↔ I apply h rw [mul_assoc, ← HAB, Units.inv_mul_cancel_left] -theorem irreducible_isUnit_mul {a b : α} (h : IsUnit a) : Irreducible (a * b) ↔ Irreducible b := +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 : αˣ) (b : α) : Irreducible (b * ↑a) ↔ Irreducible b := by +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] @@ -276,11 +276,11 @@ theorem irreducible_mul_units (a : αˣ) (b : α) : Irreducible (b * ↑a) ↔ I apply h rw [← mul_assoc, ← HAB, Units.mul_inv_cancel_right] -theorem irreducible_mul_isUnit {a b : α} (h : IsUnit a) : Irreducible (b * a) ↔ Irreducible b := +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 : α} : +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 @@ -294,7 +294,7 @@ end section CommMonoid -variable [CommMonoid α] {a : α} +variable [CommMonoid M] {a : M} theorem Irreducible.not_square (ha : Irreducible a) : ¬IsSquare a := by rw [isSquare_iff_exists_sq] @@ -307,22 +307,22 @@ end CommMonoid section CommMonoidWithZero -variable [CommMonoidWithZero α] +variable [CommMonoidWithZero M] -theorem Irreducible.prime_of_isPrimal {a : α} +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 α] {a : α} (irr : Irreducible a) : Prime a := +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 α] {a p : α} +variable [CancelCommMonoidWithZero M] {a p : M} protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := ⟨hp.not_unit, fun a b ↦ by @@ -333,10 +333,10 @@ protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := (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 α] {a : α} : Irreducible a ↔ Prime a := +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 : α} {k l : ℕ} : +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 @@ -359,8 +359,8 @@ 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 α] (x y : α) : Prop := - ∃ u : αˣ, x * u = y +def Associated [Monoid M] (x y : M) : Prop := + ∃ u : Mˣ, x * u = y /-- Notation for two elements of a monoid are associated, i.e. if one of them is another one multiplied by a unit on the right. -/ @@ -369,35 +369,35 @@ local infixl:50 " ~ᵤ " => Associated namespace Associated @[refl] -protected theorem refl [Monoid α] (x : α) : x ~ᵤ x := +protected theorem refl [Monoid M] (x : M) : x ~ᵤ x := ⟨1, by simp⟩ -protected theorem rfl [Monoid α] {x : α} : x ~ᵤ x := +protected theorem rfl [Monoid M] {x : M} : x ~ᵤ x := .refl x -instance [Monoid α] : IsRefl α Associated := +instance [Monoid M] : IsRefl M Associated := ⟨Associated.refl⟩ @[symm] -protected theorem symm [Monoid α] : ∀ {x y : α}, x ~ᵤ y → y ~ᵤ x +protected theorem symm [Monoid M] : ∀ {x y : M}, x ~ᵤ y → y ~ᵤ x | x, _, ⟨u, rfl⟩ => ⟨u⁻¹, by rw [mul_assoc, Units.mul_inv, mul_one]⟩ -instance [Monoid α] : IsSymm α Associated := +instance [Monoid M] : IsSymm M Associated := ⟨fun _ _ => Associated.symm⟩ -protected theorem comm [Monoid α] {x y : α} : x ~ᵤ y ↔ y ~ᵤ x := +protected theorem comm [Monoid M] {x y : M} : x ~ᵤ y ↔ y ~ᵤ x := ⟨Associated.symm, Associated.symm⟩ @[trans] -protected theorem trans [Monoid α] : ∀ {x y z : α}, x ~ᵤ y → y ~ᵤ z → x ~ᵤ z +protected theorem trans [Monoid M] : ∀ {x y z : M}, x ~ᵤ y → y ~ᵤ z → x ~ᵤ z | x, _, _, ⟨u, rfl⟩, ⟨v, rfl⟩ => ⟨u * v, by rw [Units.val_mul, mul_assoc]⟩ -instance [Monoid α] : IsTrans α Associated := +instance [Monoid M] : IsTrans M Associated := ⟨fun _ _ _ => Associated.trans⟩ /-- The setoid of the relation `x ~ᵤ y` iff there is a unit `u` such that `x * u = y` -/ -protected def setoid (α : Type*) [Monoid α] : - Setoid α where +protected def setoid (M : Type*) [Monoid M] : + Setoid M where r := Associated iseqv := ⟨Associated.refl, Associated.symm, Associated.trans⟩ @@ -410,11 +410,11 @@ end Associated attribute [local instance] Associated.setoid -theorem unit_associated_one [Monoid α] {u : αˣ} : (u : α) ~ᵤ 1 := +theorem unit_associated_one [Monoid M] {u : Mˣ} : (u : M) ~ᵤ 1 := ⟨u⁻¹, Units.mul_inv u⟩ @[simp] -theorem associated_one_iff_isUnit [Monoid α] {a : α} : (a : α) ~ᵤ 1 ↔ IsUnit a := +theorem associated_one_iff_isUnit [Monoid M] {a : M} : (a : M) ~ᵤ 1 ↔ IsUnit a := Iff.intro (fun h => let ⟨c, h⟩ := h.symm @@ -422,98 +422,98 @@ theorem associated_one_iff_isUnit [Monoid α] {a : α} : (a : α) ~ᵤ 1 ↔ IsU fun ⟨c, h⟩ => Associated.symm ⟨c, by simp [h]⟩ @[simp] -theorem associated_zero_iff_eq_zero [MonoidWithZero α] (a : α) : a ~ᵤ 0 ↔ a = 0 := +theorem associated_zero_iff_eq_zero [MonoidWithZero M] (a : M) : a ~ᵤ 0 ↔ a = 0 := Iff.intro (fun h => by let ⟨u, h⟩ := h.symm simpa using h.symm) fun h => h ▸ Associated.refl a -theorem associated_one_of_mul_eq_one [CommMonoid α] {a : α} (b : α) (hab : a * b = 1) : a ~ᵤ 1 := - show (Units.mkOfMulEqOne a b hab : α) ~ᵤ 1 from unit_associated_one +theorem associated_one_of_mul_eq_one [CommMonoid M] {a : M} (b : M) (hab : a * b = 1) : a ~ᵤ 1 := + show (Units.mkOfMulEqOne a b hab : M) ~ᵤ 1 from unit_associated_one -theorem associated_one_of_associated_mul_one [CommMonoid α] {a b : α} : a * b ~ᵤ 1 → a ~ᵤ 1 +theorem associated_one_of_associated_mul_one [CommMonoid M] {a b : M} : a * b ~ᵤ 1 → a ~ᵤ 1 | ⟨u, h⟩ => associated_one_of_mul_eq_one (b * u) <| by simpa [mul_assoc] using h -theorem associated_mul_unit_left {β : Type*} [Monoid β] (a u : β) (hu : IsUnit u) : +theorem associated_mul_unit_left {N : Type*} [Monoid N] (a u : N) (hu : IsUnit u) : Associated (a * u) a := let ⟨u', hu⟩ := hu ⟨u'⁻¹, hu ▸ Units.mul_inv_cancel_right _ _⟩ -theorem associated_unit_mul_left {β : Type*} [CommMonoid β] (a u : β) (hu : IsUnit u) : +theorem associated_unit_mul_left {N : Type*} [CommMonoid N] (a u : N) (hu : IsUnit u) : Associated (u * a) a := by rw [mul_comm] exact associated_mul_unit_left _ _ hu -theorem associated_mul_unit_right {β : Type*} [Monoid β] (a u : β) (hu : IsUnit u) : +theorem associated_mul_unit_right {N : Type*} [Monoid N] (a u : N) (hu : IsUnit u) : Associated a (a * u) := (associated_mul_unit_left a u hu).symm -theorem associated_unit_mul_right {β : Type*} [CommMonoid β] (a u : β) (hu : IsUnit u) : +theorem associated_unit_mul_right {N : Type*} [CommMonoid N] (a u : N) (hu : IsUnit u) : Associated a (u * a) := (associated_unit_mul_left a u hu).symm -theorem associated_mul_isUnit_left_iff {β : Type*} [Monoid β] {a u b : β} (hu : IsUnit u) : +theorem associated_mul_isUnit_left_iff {N : Type*} [Monoid N] {a u b : N} (hu : IsUnit u) : Associated (a * u) b ↔ Associated a b := ⟨(associated_mul_unit_right _ _ hu).trans, (associated_mul_unit_left _ _ hu).trans⟩ -theorem associated_isUnit_mul_left_iff {β : Type*} [CommMonoid β] {u a b : β} (hu : IsUnit u) : +theorem associated_isUnit_mul_left_iff {N : Type*} [CommMonoid N] {u a b : N} (hu : IsUnit u) : Associated (u * a) b ↔ Associated a b := by rw [mul_comm] exact associated_mul_isUnit_left_iff hu -theorem associated_mul_isUnit_right_iff {β : Type*} [Monoid β] {a b u : β} (hu : IsUnit u) : +theorem associated_mul_isUnit_right_iff {N : Type*} [Monoid N] {a b u : N} (hu : IsUnit u) : Associated a (b * u) ↔ Associated a b := Associated.comm.trans <| (associated_mul_isUnit_left_iff hu).trans Associated.comm -theorem associated_isUnit_mul_right_iff {β : Type*} [CommMonoid β] {a u b : β} (hu : IsUnit u) : +theorem associated_isUnit_mul_right_iff {N : Type*} [CommMonoid N] {a u b : N} (hu : IsUnit u) : Associated a (u * b) ↔ Associated a b := Associated.comm.trans <| (associated_isUnit_mul_left_iff hu).trans Associated.comm @[simp] -theorem associated_mul_unit_left_iff {β : Type*} [Monoid β] {a b : β} {u : Units β} : +theorem associated_mul_unit_left_iff {N : Type*} [Monoid N] {a b : N} {u : Units N} : Associated (a * u) b ↔ Associated a b := associated_mul_isUnit_left_iff u.isUnit @[simp] -theorem associated_unit_mul_left_iff {β : Type*} [CommMonoid β] {a b : β} {u : Units β} : +theorem associated_unit_mul_left_iff {N : Type*} [CommMonoid N] {a b : N} {u : Units N} : Associated (↑u * a) b ↔ Associated a b := associated_isUnit_mul_left_iff u.isUnit @[simp] -theorem associated_mul_unit_right_iff {β : Type*} [Monoid β] {a b : β} {u : Units β} : +theorem associated_mul_unit_right_iff {N : Type*} [Monoid N] {a b : N} {u : Units N} : Associated a (b * u) ↔ Associated a b := associated_mul_isUnit_right_iff u.isUnit @[simp] -theorem associated_unit_mul_right_iff {β : Type*} [CommMonoid β] {a b : β} {u : Units β} : +theorem associated_unit_mul_right_iff {N : Type*} [CommMonoid N] {a b : N} {u : Units N} : Associated a (↑u * b) ↔ Associated a b := associated_isUnit_mul_right_iff u.isUnit -theorem Associated.mul_left [Monoid α] (a : α) {b c : α} (h : b ~ᵤ c) : a * b ~ᵤ a * c := by +theorem Associated.mul_left [Monoid M] (a : M) {b c : M} (h : b ~ᵤ c) : a * b ~ᵤ a * c := by obtain ⟨d, rfl⟩ := h; exact ⟨d, mul_assoc _ _ _⟩ -theorem Associated.mul_right [CommMonoid α] {a b : α} (h : a ~ᵤ b) (c : α) : a * c ~ᵤ b * c := by +theorem Associated.mul_right [CommMonoid M] {a b : M} (h : a ~ᵤ b) (c : M) : a * c ~ᵤ b * c := by obtain ⟨d, rfl⟩ := h; exact ⟨d, mul_right_comm _ _ _⟩ -theorem Associated.mul_mul [CommMonoid α] {a₁ a₂ b₁ b₂ : α} +theorem Associated.mul_mul [CommMonoid M] {a₁ a₂ b₁ b₂ : M} (h₁ : a₁ ~ᵤ b₁) (h₂ : a₂ ~ᵤ b₂) : a₁ * a₂ ~ᵤ b₁ * b₂ := (h₁.mul_right _).trans (h₂.mul_left _) -theorem Associated.pow_pow [CommMonoid α] {a b : α} {n : ℕ} (h : a ~ᵤ b) : a ^ n ~ᵤ b ^ n := by +theorem Associated.pow_pow [CommMonoid M] {a b : M} {n : ℕ} (h : a ~ᵤ b) : a ^ n ~ᵤ b ^ n := by induction n with | zero => simp [Associated.refl] | succ n ih => convert h.mul_mul ih <;> rw [pow_succ'] -protected theorem Associated.dvd [Monoid α] {a b : α} : a ~ᵤ b → a ∣ b := fun ⟨u, hu⟩ => +protected theorem Associated.dvd [Monoid M] {a b : M} : a ~ᵤ b → a ∣ b := fun ⟨u, hu⟩ => ⟨u, hu.symm⟩ -protected theorem Associated.dvd' [Monoid α] {a b : α} (h : a ~ᵤ b) : b ∣ a := +protected theorem Associated.dvd' [Monoid M] {a b : M} (h : a ~ᵤ b) : b ∣ a := h.symm.dvd -protected theorem Associated.dvd_dvd [Monoid α] {a b : α} (h : a ~ᵤ b) : a ∣ b ∧ b ∣ a := +protected theorem Associated.dvd_dvd [Monoid M] {a b : M} (h : a ~ᵤ b) : a ∣ b ∧ b ∣ a := ⟨h.dvd, h.symm.dvd⟩ -theorem associated_of_dvd_dvd [CancelMonoidWithZero α] {a b : α} (hab : a ∣ b) (hba : b ∣ a) : +theorem associated_of_dvd_dvd [CancelMonoidWithZero M] {a b : M} (hab : a ∣ b) (hba : b ∣ a) : a ~ᵤ b := by rcases hab with ⟨c, rfl⟩ rcases hba with ⟨d, a_eq⟩ @@ -529,40 +529,40 @@ theorem associated_of_dvd_dvd [CancelMonoidWithZero α] {a b : α} (hab : a ∣ have hdc : d * c = 1 := mul_left_cancel₀ hac0 this exact ⟨⟨c, d, hcd, hdc⟩, rfl⟩ -theorem dvd_dvd_iff_associated [CancelMonoidWithZero α] {a b : α} : a ∣ b ∧ b ∣ a ↔ a ~ᵤ b := +theorem dvd_dvd_iff_associated [CancelMonoidWithZero M] {a b : M} : a ∣ b ∧ b ∣ a ↔ a ~ᵤ b := ⟨fun ⟨h1, h2⟩ => associated_of_dvd_dvd h1 h2, Associated.dvd_dvd⟩ -instance [CancelMonoidWithZero α] [DecidableRel ((· ∣ ·) : α → α → Prop)] : - DecidableRel ((· ~ᵤ ·) : α → α → Prop) := fun _ _ => decidable_of_iff _ dvd_dvd_iff_associated +instance [CancelMonoidWithZero M] [DecidableRel ((· ∣ ·) : M → M → Prop)] : + DecidableRel ((· ~ᵤ ·) : M → M → Prop) := fun _ _ => decidable_of_iff _ dvd_dvd_iff_associated -theorem Associated.dvd_iff_dvd_left [Monoid α] {a b c : α} (h : a ~ᵤ b) : a ∣ c ↔ b ∣ c := +theorem Associated.dvd_iff_dvd_left [Monoid M] {a b c : M} (h : a ~ᵤ b) : a ∣ c ↔ b ∣ c := let ⟨_, hu⟩ := h hu ▸ Units.mul_right_dvd.symm -theorem Associated.dvd_iff_dvd_right [Monoid α] {a b c : α} (h : b ~ᵤ c) : a ∣ b ↔ a ∣ c := +theorem Associated.dvd_iff_dvd_right [Monoid M] {a b c : M} (h : b ~ᵤ c) : a ∣ b ↔ a ∣ c := let ⟨_, hu⟩ := h hu ▸ Units.dvd_mul_right.symm -theorem Associated.eq_zero_iff [MonoidWithZero α] {a b : α} (h : a ~ᵤ b) : a = 0 ↔ b = 0 := by +theorem Associated.eq_zero_iff [MonoidWithZero M] {a b : M} (h : a ~ᵤ b) : a = 0 ↔ b = 0 := by obtain ⟨u, rfl⟩ := h rw [← Units.eq_mul_inv_iff_mul_eq, zero_mul] -theorem Associated.ne_zero_iff [MonoidWithZero α] {a b : α} (h : a ~ᵤ b) : a ≠ 0 ↔ b ≠ 0 := +theorem Associated.ne_zero_iff [MonoidWithZero M] {a b : M} (h : a ~ᵤ b) : a ≠ 0 ↔ b ≠ 0 := not_congr h.eq_zero_iff -theorem Associated.neg_left [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : +theorem Associated.neg_left [Monoid M] [HasDistribNeg M] {a b : M} (h : Associated a b) : Associated (-a) b := let ⟨u, hu⟩ := h; ⟨-u, by simp [hu]⟩ -theorem Associated.neg_right [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : +theorem Associated.neg_right [Monoid M] [HasDistribNeg M] {a b : M} (h : Associated a b) : Associated a (-b) := h.symm.neg_left.symm -theorem Associated.neg_neg [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : +theorem Associated.neg_neg [Monoid M] [HasDistribNeg M] {a b : M} (h : Associated a b) : Associated (-a) (-b) := h.neg_left.neg_right -protected theorem Associated.prime [CommMonoidWithZero α] {p q : α} (h : p ~ᵤ q) (hp : Prime p) : +protected theorem Associated.prime [CommMonoidWithZero M] {p q : M} (h : p ~ᵤ q) (hp : Prime p) : Prime q := ⟨h.ne_zero_iff.1 hp.ne_zero, let ⟨u, hu⟩ := h @@ -572,7 +572,7 @@ protected theorem Associated.prime [CommMonoidWithZero α] {p q : α} (h : p ~ intro a b exact hp.dvd_or_dvd⟩⟩ -theorem prime_mul_iff [CancelCommMonoidWithZero α] {x y : α} : +theorem prime_mul_iff [CancelCommMonoidWithZero M] {x y : M} : Prime (x * y) ↔ (Prime x ∧ IsUnit y) ∨ (IsUnit x ∧ Prime y) := by refine ⟨fun h ↦ ?_, ?_⟩ · rcases of_irreducible_mul h.irreducible with hx | hy @@ -583,7 +583,7 @@ theorem prime_mul_iff [CancelCommMonoidWithZero α] {x y : α} : · exact (associated_unit_mul_right y x hx).prime hy @[simp] -lemma prime_pow_iff [CancelCommMonoidWithZero α] {p : α} {n : ℕ} : +lemma prime_pow_iff [CancelCommMonoidWithZero M] {p : M} {n : ℕ} : Prime (p ^ n) ↔ Prime p ∧ n = 1 := by refine ⟨fun hp ↦ ?_, fun ⟨hp, hn⟩ ↦ by simpa [hn]⟩ suffices n = 1 by aesop @@ -598,7 +598,7 @@ lemma prime_pow_iff [CancelCommMonoidWithZero α] {p : α} {n : ℕ} : · exfalso exact hpn.not_unit (hp.pow n) -theorem Irreducible.dvd_iff [Monoid α] {x y : α} (hx : Irreducible x) : +theorem Irreducible.dvd_iff [Monoid M] {x y : M} (hx : Irreducible x) : y ∣ x ↔ IsUnit y ∨ Associated x y := by constructor · rintro ⟨z, hz⟩ @@ -610,67 +610,67 @@ theorem Irreducible.dvd_iff [Monoid α] {x y : α} (hx : Irreducible x) : · exact hy.dvd · exact h.symm.dvd -theorem Irreducible.associated_of_dvd [Monoid α] {p q : α} (p_irr : Irreducible p) +theorem Irreducible.associated_of_dvd [Monoid M] {p q : M} (p_irr : Irreducible p) (q_irr : Irreducible q) (dvd : p ∣ q) : Associated p q := ((q_irr.dvd_iff.mp dvd).resolve_left p_irr.not_unit).symm -theorem Irreducible.dvd_irreducible_iff_associated [Monoid α] {p q : α} +theorem Irreducible.dvd_irreducible_iff_associated [Monoid M] {p q : M} (pp : Irreducible p) (qp : Irreducible q) : p ∣ q ↔ Associated p q := ⟨Irreducible.associated_of_dvd pp qp, Associated.dvd⟩ -theorem Prime.associated_of_dvd [CancelCommMonoidWithZero α] {p q : α} (p_prime : Prime p) +theorem Prime.associated_of_dvd [CancelCommMonoidWithZero M] {p q : M} (p_prime : Prime p) (q_prime : Prime q) (dvd : p ∣ q) : Associated p q := p_prime.irreducible.associated_of_dvd q_prime.irreducible dvd -theorem Prime.dvd_prime_iff_associated [CancelCommMonoidWithZero α] {p q : α} (pp : Prime p) +theorem Prime.dvd_prime_iff_associated [CancelCommMonoidWithZero M] {p q : M} (pp : Prime p) (qp : Prime q) : p ∣ q ↔ Associated p q := pp.irreducible.dvd_irreducible_iff_associated qp.irreducible -theorem Associated.prime_iff [CommMonoidWithZero α] {p q : α} (h : p ~ᵤ q) : Prime p ↔ Prime q := +theorem Associated.prime_iff [CommMonoidWithZero M] {p q : M} (h : p ~ᵤ q) : Prime p ↔ Prime q := ⟨h.prime, h.symm.prime⟩ -protected theorem Associated.isUnit [Monoid α] {a b : α} (h : a ~ᵤ b) : IsUnit a → IsUnit b := +protected theorem Associated.isUnit [Monoid M] {a b : M} (h : a ~ᵤ b) : IsUnit a → IsUnit b := let ⟨u, hu⟩ := h fun ⟨v, hv⟩ => ⟨v * u, by simp [hv, hu.symm]⟩ -theorem Associated.isUnit_iff [Monoid α] {a b : α} (h : a ~ᵤ b) : IsUnit a ↔ IsUnit b := +theorem Associated.isUnit_iff [Monoid M] {a b : M} (h : a ~ᵤ b) : IsUnit a ↔ IsUnit b := ⟨h.isUnit, h.symm.isUnit⟩ -theorem Irreducible.isUnit_iff_not_associated_of_dvd [Monoid α] - {x y : α} (hx : Irreducible x) (hy : y ∣ x) : IsUnit y ↔ ¬ Associated x y := +theorem Irreducible.isUnit_iff_not_associated_of_dvd [Monoid M] + {x y : M} (hx : Irreducible x) (hy : y ∣ x) : IsUnit y ↔ ¬ Associated x y := ⟨fun hy hxy => hx.1 (hxy.symm.isUnit hy), (hx.dvd_iff.mp hy).resolve_right⟩ -protected theorem Associated.irreducible [Monoid α] {p q : α} (h : p ~ᵤ q) (hp : Irreducible p) : +protected theorem Associated.irreducible [Monoid M] {p q : M} (h : p ~ᵤ q) (hp : Irreducible p) : Irreducible q := ⟨mt h.symm.isUnit hp.1, let ⟨u, hu⟩ := h fun a b hab => - have hpab : p = a * (b * (u⁻¹ : αˣ)) := + have hpab : p = a * (b * (u⁻¹ : Mˣ)) := calc - p = p * u * (u⁻¹ : αˣ) := by simp + p = p * u * (u⁻¹ : Mˣ) := by simp _ = _ := by rw [hu]; simp [hab, mul_assoc] (hp.isUnit_or_isUnit hpab).elim Or.inl fun ⟨v, hv⟩ => Or.inr ⟨v * u, by simp [hv]⟩⟩ -protected theorem Associated.irreducible_iff [Monoid α] {p q : α} (h : p ~ᵤ q) : +protected theorem Associated.irreducible_iff [Monoid M] {p q : M} (h : p ~ᵤ q) : Irreducible p ↔ Irreducible q := ⟨h.irreducible, h.symm.irreducible⟩ -theorem Associated.of_mul_left [CancelCommMonoidWithZero α] {a b c d : α} (h : a * b ~ᵤ c * d) +theorem Associated.of_mul_left [CancelCommMonoidWithZero M] {a b c d : M} (h : a * b ~ᵤ c * d) (h₁ : a ~ᵤ c) (ha : a ≠ 0) : b ~ᵤ d := let ⟨u, hu⟩ := h let ⟨v, hv⟩ := Associated.symm h₁ - ⟨u * (v : αˣ), + ⟨u * (v : Mˣ), mul_left_cancel₀ ha (by - rw [← hv, mul_assoc c (v : α) d, mul_left_comm c, ← hu] + rw [← hv, mul_assoc c (v : M) d, mul_left_comm c, ← hu] simp [hv.symm, mul_assoc, mul_comm, mul_left_comm])⟩ -theorem Associated.of_mul_right [CancelCommMonoidWithZero α] {a b c d : α} : +theorem Associated.of_mul_right [CancelCommMonoidWithZero M] {a b c d : M} : a * b ~ᵤ c * d → b ~ᵤ d → b ≠ 0 → a ~ᵤ c := by rw [mul_comm a, mul_comm c]; exact Associated.of_mul_left -theorem Associated.of_pow_associated_of_prime [CancelCommMonoidWithZero α] {p₁ p₂ : α} {k₁ k₂ : ℕ} +theorem Associated.of_pow_associated_of_prime [CancelCommMonoidWithZero M] {p₁ p₂ : M} {k₁ k₂ : ℕ} (hp₁ : Prime p₁) (hp₂ : Prime p₂) (hk₁ : 0 < k₁) (h : p₁ ^ k₁ ~ᵤ p₂ ^ k₂) : p₁ ~ᵤ p₂ := by have : p₁ ∣ p₂ ^ k₂ := by rw [← h.dvd_iff_dvd_right] @@ -678,33 +678,33 @@ theorem Associated.of_pow_associated_of_prime [CancelCommMonoidWithZero α] {p rw [← hp₁.dvd_prime_iff_associated hp₂] exact hp₁.dvd_of_dvd_pow this -theorem Associated.of_pow_associated_of_prime' [CancelCommMonoidWithZero α] {p₁ p₂ : α} {k₁ k₂ : ℕ} +theorem Associated.of_pow_associated_of_prime' [CancelCommMonoidWithZero M] {p₁ p₂ : M} {k₁ k₂ : ℕ} (hp₁ : Prime p₁) (hp₂ : Prime p₂) (hk₂ : 0 < k₂) (h : p₁ ^ k₁ ~ᵤ p₂ ^ k₂) : p₁ ~ᵤ p₂ := (h.symm.of_pow_associated_of_prime hp₂ hp₁ hk₂).symm /-- See also `Irreducible.coprime_iff_not_dvd`. -/ -lemma Irreducible.isRelPrime_iff_not_dvd [Monoid α] {p n : α} (hp : Irreducible p) : +lemma Irreducible.isRelPrime_iff_not_dvd [Monoid M] {p n : M} (hp : Irreducible p) : IsRelPrime p n ↔ ¬ p ∣ n := by refine ⟨fun h contra ↦ hp.not_unit (h dvd_rfl contra), fun hpn d hdp hdn ↦ ?_⟩ contrapose! hpn suffices Associated p d from this.dvd.trans hdn exact (hp.dvd_iff.mp hdp).resolve_left hpn -lemma Irreducible.dvd_or_isRelPrime [Monoid α] {p n : α} (hp : Irreducible p) : +lemma Irreducible.dvd_or_isRelPrime [Monoid M] {p n : M} (hp : Irreducible p) : p ∣ n ∨ IsRelPrime p n := Classical.or_iff_not_imp_left.mpr hp.isRelPrime_iff_not_dvd.2 section UniqueUnits -variable [Monoid α] [Unique αˣ] +variable [Monoid M] [Unique Mˣ] -theorem associated_iff_eq {x y : α} : x ~ᵤ y ↔ x = y := by +theorem associated_iff_eq {x y : M} : x ~ᵤ y ↔ x = y := by constructor · rintro ⟨c, rfl⟩ rw [units_eq_one c, Units.val_one, mul_one] · rintro rfl rfl -theorem associated_eq_eq : (Associated : α → α → Prop) = Eq := by +theorem associated_eq_eq : (Associated : M → M → Prop) = Eq := by ext rw [associated_iff_eq] @@ -732,86 +732,86 @@ end UniqueUnits₀ /-- The quotient of a monoid by the `Associated` relation. Two elements `x` and `y` are associated iff there is a unit `u` such that `x * u = y`. There is a natural - monoid structure on `Associates α`. -/ -abbrev Associates (α : Type*) [Monoid α] : Type _ := - Quotient (Associated.setoid α) + monoid structure on `Associates M`. -/ +abbrev Associates (M : Type*) [Monoid M] : Type _ := + Quotient (Associated.setoid M) namespace Associates open Associated -/-- The canonical quotient map from a monoid `α` into the `Associates` of `α` -/ -protected abbrev mk {α : Type*} [Monoid α] (a : α) : Associates α := +/-- The canonical quotient map from a monoid `M` into the `Associates` of `M` -/ +protected abbrev mk {M : Type*} [Monoid M] (a : M) : Associates M := ⟦a⟧ -instance [Monoid α] : Inhabited (Associates α) := +instance [Monoid M] : Inhabited (Associates M) := ⟨⟦1⟧⟩ -theorem mk_eq_mk_iff_associated [Monoid α] {a b : α} : Associates.mk a = Associates.mk b ↔ a ~ᵤ b := +theorem mk_eq_mk_iff_associated [Monoid M] {a b : M} : Associates.mk a = Associates.mk b ↔ a ~ᵤ b := Iff.intro Quotient.exact Quot.sound -theorem quotient_mk_eq_mk [Monoid α] (a : α) : ⟦a⟧ = Associates.mk a := +theorem quotient_mk_eq_mk [Monoid M] (a : M) : ⟦a⟧ = Associates.mk a := rfl -theorem quot_mk_eq_mk [Monoid α] (a : α) : Quot.mk Setoid.r a = Associates.mk a := +theorem quot_mk_eq_mk [Monoid M] (a : M) : Quot.mk Setoid.r a = Associates.mk a := rfl @[simp] -theorem quot_out [Monoid α] (a : Associates α) : Associates.mk (Quot.out a) = a := by +theorem quot_out [Monoid M] (a : Associates M) : Associates.mk (Quot.out a) = a := by rw [← quot_mk_eq_mk, Quot.out_eq] -theorem mk_quot_out [Monoid α] (a : α) : Quot.out (Associates.mk a) ~ᵤ a := by +theorem mk_quot_out [Monoid M] (a : M) : Quot.out (Associates.mk a) ~ᵤ a := by rw [← Associates.mk_eq_mk_iff_associated, Associates.quot_out] -theorem forall_associated [Monoid α] {p : Associates α → Prop} : +theorem forall_associated [Monoid M] {p : Associates M → Prop} : (∀ a, p a) ↔ ∀ a, p (Associates.mk a) := Iff.intro (fun h _ => h _) fun h a => Quotient.inductionOn a h -theorem mk_surjective [Monoid α] : Function.Surjective (@Associates.mk α _) := +theorem mk_surjective [Monoid M] : Function.Surjective (@Associates.mk M _) := forall_associated.2 fun a => ⟨a, rfl⟩ -instance [Monoid α] : One (Associates α) := +instance [Monoid M] : One (Associates M) := ⟨⟦1⟧⟩ @[simp] -theorem mk_one [Monoid α] : Associates.mk (1 : α) = 1 := +theorem mk_one [Monoid M] : Associates.mk (1 : M) = 1 := rfl -theorem one_eq_mk_one [Monoid α] : (1 : Associates α) = Associates.mk 1 := +theorem one_eq_mk_one [Monoid M] : (1 : Associates M) = Associates.mk 1 := rfl @[simp] -theorem mk_eq_one [Monoid α] {a : α} : Associates.mk a = 1 ↔ IsUnit a := by +theorem mk_eq_one [Monoid M] {a : M} : Associates.mk a = 1 ↔ IsUnit a := by rw [← mk_one, mk_eq_mk_iff_associated, associated_one_iff_isUnit] -instance [Monoid α] : Bot (Associates α) := +instance [Monoid M] : Bot (Associates M) := ⟨1⟩ -theorem bot_eq_one [Monoid α] : (⊥ : Associates α) = 1 := +theorem bot_eq_one [Monoid M] : (⊥ : Associates M) = 1 := rfl -theorem exists_rep [Monoid α] (a : Associates α) : ∃ a0 : α, Associates.mk a0 = a := +theorem exists_rep [Monoid M] (a : Associates M) : ∃ a0 : M, Associates.mk a0 = a := Quot.exists_rep a -instance [Monoid α] [Subsingleton α] : - Unique (Associates α) where +instance [Monoid M] [Subsingleton M] : + Unique (Associates M) where default := 1 uniq := forall_associated.2 fun _ ↦ mk_eq_one.2 <| isUnit_of_subsingleton _ -theorem mk_injective [Monoid α] [Unique (Units α)] : Function.Injective (@Associates.mk α _) := +theorem mk_injective [Monoid M] [Unique (Units M)] : Function.Injective (@Associates.mk M _) := fun _ _ h => associated_iff_eq.mp (Associates.mk_eq_mk_iff_associated.mp h) section CommMonoid -variable [CommMonoid α] +variable [CommMonoid M] -instance instMul : Mul (Associates α) := +instance instMul : Mul (Associates M) := ⟨Quotient.map₂ (· * ·) fun _ _ h₁ _ _ h₂ ↦ h₁.mul_mul h₂⟩ -theorem mk_mul_mk {x y : α} : Associates.mk x * Associates.mk y = Associates.mk (x * y) := +theorem mk_mul_mk {x y : M} : Associates.mk x * Associates.mk y = Associates.mk (x * y) := rfl -instance instCommMonoid : CommMonoid (Associates α) where +instance instCommMonoid : CommMonoid (Associates M) where one := 1 mul := (· * ·) mul_one a' := Quotient.inductionOn a' fun a => show ⟦a * 1⟧ = ⟦a⟧ by simp @@ -822,32 +822,32 @@ instance instCommMonoid : CommMonoid (Associates α) where mul_comm a' b' := Quotient.inductionOn₂ a' b' fun a b => show ⟦a * b⟧ = ⟦b * a⟧ by rw [mul_comm] -instance instPreorder : Preorder (Associates α) where +instance instPreorder : Preorder (Associates M) where le := Dvd.dvd le_refl := dvd_refl le_trans a b c := dvd_trans /-- `Associates.mk` as a `MonoidHom`. -/ -protected def mkMonoidHom : α →* Associates α where +protected def mkMonoidHom : M →* Associates M where toFun := Associates.mk map_one' := mk_one map_mul' _ _ := mk_mul_mk @[simp] -theorem mkMonoidHom_apply (a : α) : Associates.mkMonoidHom a = Associates.mk a := +theorem mkMonoidHom_apply (a : M) : Associates.mkMonoidHom a = Associates.mk a := rfl -theorem associated_map_mk {f : Associates α →* α} (hinv : Function.RightInverse f Associates.mk) - (a : α) : a ~ᵤ f (Associates.mk a) := +theorem associated_map_mk {f : Associates M →* M} (hinv : Function.RightInverse f Associates.mk) + (a : M) : a ~ᵤ f (Associates.mk a) := Associates.mk_eq_mk_iff_associated.1 (hinv (Associates.mk a)).symm -theorem mk_pow (a : α) (n : ℕ) : Associates.mk (a ^ n) = Associates.mk a ^ n := by +theorem mk_pow (a : M) (n : ℕ) : Associates.mk (a ^ n) = Associates.mk a ^ n := by induction n <;> simp [*, pow_succ, Associates.mk_mul_mk.symm] -theorem dvd_eq_le : ((· ∣ ·) : Associates α → Associates α → Prop) = (· ≤ ·) := +theorem dvd_eq_le : ((· ∣ ·) : Associates M → Associates M → Prop) = (· ≤ ·) := rfl -instance uniqueUnits : Unique (Associates α)ˣ where +instance uniqueUnits : Unique (Associates M)ˣ where uniq := by rintro ⟨a, b, hab, hba⟩ revert hab hba @@ -858,16 +858,16 @@ instance uniqueUnits : Unique (Associates α)ˣ where @[deprecated (since := "2024-07-22")] protected alias units_eq_one := Subsingleton.elim @[simp] -theorem coe_unit_eq_one (u : (Associates α)ˣ) : (u : Associates α) = 1 := by +theorem coe_unit_eq_one (u : (Associates M)ˣ) : (u : Associates M) = 1 := by simp [eq_iff_true_of_subsingleton] -theorem isUnit_iff_eq_one (a : Associates α) : IsUnit a ↔ a = 1 := +theorem isUnit_iff_eq_one (a : Associates M) : IsUnit a ↔ a = 1 := Iff.intro (fun ⟨_, h⟩ => h ▸ coe_unit_eq_one _) fun h => h.symm ▸ isUnit_one -theorem isUnit_iff_eq_bot {a : Associates α} : IsUnit a ↔ a = ⊥ := by +theorem isUnit_iff_eq_bot {a : Associates M} : IsUnit a ↔ a = ⊥ := by rw [Associates.isUnit_iff_eq_one, bot_eq_one] -theorem isUnit_mk {a : α} : IsUnit (Associates.mk a) ↔ IsUnit a := +theorem isUnit_mk {a : M} : IsUnit (Associates.mk a) ↔ IsUnit a := calc IsUnit (Associates.mk a) ↔ a ~ᵤ 1 := by rw [isUnit_iff_eq_one, one_eq_mk_one, mk_eq_mk_iff_associated] @@ -875,27 +875,27 @@ theorem isUnit_mk {a : α} : IsUnit (Associates.mk a) ↔ IsUnit a := section Order -theorem mul_mono {a b c d : Associates α} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := +theorem mul_mono {a b c d : Associates M} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := let ⟨x, hx⟩ := h₁ let ⟨y, hy⟩ := h₂ ⟨x * y, by simp [hx, hy, mul_comm, mul_assoc, mul_left_comm]⟩ -theorem one_le {a : Associates α} : 1 ≤ a := +theorem one_le {a : Associates M} : 1 ≤ a := Dvd.intro _ (one_mul a) -theorem le_mul_right {a b : Associates α} : a ≤ a * b := +theorem le_mul_right {a b : Associates M} : a ≤ a * b := ⟨b, rfl⟩ -theorem le_mul_left {a b : Associates α} : a ≤ b * a := by rw [mul_comm]; exact le_mul_right +theorem le_mul_left {a b : Associates M} : a ≤ b * a := by rw [mul_comm]; exact le_mul_right -instance instOrderBot : OrderBot (Associates α) where +instance instOrderBot : OrderBot (Associates M) where bot := 1 bot_le _ := one_le end Order @[simp] -theorem mk_dvd_mk {a b : α} : Associates.mk a ∣ Associates.mk b ↔ a ∣ b := by +theorem mk_dvd_mk {a b : M} : Associates.mk a ∣ Associates.mk b ↔ a ∣ b := by simp only [dvd_def, mk_surjective.exists, mk_mul_mk, mk_eq_mk_iff_associated, Associated.comm (x := b)] constructor @@ -904,18 +904,18 @@ theorem mk_dvd_mk {a b : α} : Associates.mk a ∣ Associates.mk b ↔ a ∣ b : · rintro ⟨c, rfl⟩ use c -theorem dvd_of_mk_le_mk {a b : α} : Associates.mk a ≤ Associates.mk b → a ∣ b := +theorem dvd_of_mk_le_mk {a b : M} : Associates.mk a ≤ Associates.mk b → a ∣ b := mk_dvd_mk.mp -theorem mk_le_mk_of_dvd {a b : α} : a ∣ b → Associates.mk a ≤ Associates.mk b := +theorem mk_le_mk_of_dvd {a b : M} : a ∣ b → Associates.mk a ≤ Associates.mk b := mk_dvd_mk.mpr -theorem mk_le_mk_iff_dvd {a b : α} : Associates.mk a ≤ Associates.mk b ↔ a ∣ b := mk_dvd_mk +theorem mk_le_mk_iff_dvd {a b : M} : Associates.mk a ≤ Associates.mk b ↔ a ∣ b := mk_dvd_mk @[deprecated (since := "2024-03-16")] alias mk_le_mk_iff_dvd_iff := mk_le_mk_iff_dvd @[simp] -theorem isPrimal_mk {a : α} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by +theorem isPrimal_mk {a : M} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by simp_rw [IsPrimal, forall_associated, mk_surjective.exists, mk_mul_mk, mk_dvd_mk] constructor <;> intro h b c dvd <;> obtain ⟨a₁, a₂, h₁, h₂, eq⟩ := @h b c dvd · obtain ⟨u, rfl⟩ := mk_eq_mk_iff_associated.mp eq.symm @@ -925,80 +925,80 @@ theorem isPrimal_mk {a : α} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by @[deprecated (since := "2024-03-16")] alias isPrimal_iff := isPrimal_mk @[simp] -theorem decompositionMonoid_iff : DecompositionMonoid (Associates α) ↔ DecompositionMonoid α := by +theorem decompositionMonoid_iff : DecompositionMonoid (Associates M) ↔ DecompositionMonoid M := by simp_rw [_root_.decompositionMonoid_iff, forall_associated, isPrimal_mk] -instance instDecompositionMonoid [DecompositionMonoid α] : DecompositionMonoid (Associates α) := +instance instDecompositionMonoid [DecompositionMonoid M] : DecompositionMonoid (Associates M) := decompositionMonoid_iff.mpr ‹_› @[simp] -theorem mk_isRelPrime_iff {a b : α} : +theorem mk_isRelPrime_iff {a b : M} : IsRelPrime (Associates.mk a) (Associates.mk b) ↔ IsRelPrime a b := by simp_rw [IsRelPrime, forall_associated, mk_dvd_mk, isUnit_mk] end CommMonoid -instance [Zero α] [Monoid α] : Zero (Associates α) := +instance [Zero M] [Monoid M] : Zero (Associates M) := ⟨⟦0⟧⟩ -instance [Zero α] [Monoid α] : Top (Associates α) := +instance [Zero M] [Monoid M] : Top (Associates M) := ⟨0⟩ -@[simp] theorem mk_zero [Zero α] [Monoid α] : Associates.mk (0 : α) = 0 := rfl +@[simp] theorem mk_zero [Zero M] [Monoid M] : Associates.mk (0 : M) = 0 := rfl section MonoidWithZero -variable [MonoidWithZero α] +variable [MonoidWithZero M] @[simp] -theorem mk_eq_zero {a : α} : Associates.mk a = 0 ↔ a = 0 := +theorem mk_eq_zero {a : M} : Associates.mk a = 0 ↔ a = 0 := ⟨fun h => (associated_zero_iff_eq_zero a).1 <| Quotient.exact h, fun h => h.symm ▸ rfl⟩ @[simp] -theorem quot_out_zero : Quot.out (0 : Associates α) = 0 := by rw [← mk_eq_zero, quot_out] +theorem quot_out_zero : Quot.out (0 : Associates M) = 0 := by rw [← mk_eq_zero, quot_out] -theorem mk_ne_zero {a : α} : Associates.mk a ≠ 0 ↔ a ≠ 0 := +theorem mk_ne_zero {a : M} : Associates.mk a ≠ 0 ↔ a ≠ 0 := not_congr mk_eq_zero -instance [Nontrivial α] : Nontrivial (Associates α) := +instance [Nontrivial M] : Nontrivial (Associates M) := ⟨⟨1, 0, mk_ne_zero.2 one_ne_zero⟩⟩ -theorem exists_non_zero_rep {a : Associates α} : a ≠ 0 → ∃ a0 : α, a0 ≠ 0 ∧ Associates.mk a0 = a := +theorem exists_non_zero_rep {a : Associates M} : a ≠ 0 → ∃ a0 : M, a0 ≠ 0 ∧ Associates.mk a0 = a := Quotient.inductionOn a fun b nz => ⟨b, mt (congr_arg Quotient.mk'') nz, rfl⟩ end MonoidWithZero section CommMonoidWithZero -variable [CommMonoidWithZero α] +variable [CommMonoidWithZero M] -instance instCommMonoidWithZero : CommMonoidWithZero (Associates α) where +instance instCommMonoidWithZero : CommMonoidWithZero (Associates M) where zero_mul := forall_associated.2 fun a ↦ by rw [← mk_zero, mk_mul_mk, zero_mul] mul_zero := forall_associated.2 fun a ↦ by rw [← mk_zero, mk_mul_mk, mul_zero] -instance instOrderTop : OrderTop (Associates α) where +instance instOrderTop : OrderTop (Associates M) where top := 0 le_top := dvd_zero -@[simp] protected theorem le_zero (a : Associates α) : a ≤ 0 := le_top +@[simp] protected theorem le_zero (a : Associates M) : a ≤ 0 := le_top -instance instBoundedOrder : BoundedOrder (Associates α) where +instance instBoundedOrder : BoundedOrder (Associates M) where -instance [DecidableRel ((· ∣ ·) : α → α → Prop)] : - DecidableRel ((· ∣ ·) : Associates α → Associates α → Prop) := fun a b => +instance [DecidableRel ((· ∣ ·) : M → M → Prop)] : + DecidableRel ((· ∣ ·) : Associates M → Associates M → Prop) := fun a b => Quotient.recOnSubsingleton₂ a b fun _ _ => decidable_of_iff' _ mk_dvd_mk -theorem Prime.le_or_le {p : Associates α} (hp : Prime p) {a b : Associates α} (h : p ≤ a * b) : +theorem Prime.le_or_le {p : Associates M} (hp : Prime p) {a b : Associates M} (h : p ≤ a * b) : p ≤ a ∨ p ≤ b := hp.2.2 a b h @[simp] -theorem prime_mk {p : α} : Prime (Associates.mk p) ↔ Prime p := by +theorem prime_mk {p : M} : Prime (Associates.mk p) ↔ Prime p := by rw [Prime, _root_.Prime, forall_associated] simp only [forall_associated, mk_ne_zero, isUnit_mk, mk_mul_mk, mk_dvd_mk] @[simp] -theorem irreducible_mk {a : α} : Irreducible (Associates.mk a) ↔ Irreducible a := by +theorem irreducible_mk {a : M} : Irreducible (Associates.mk a) ↔ Irreducible a := by simp only [irreducible_iff, isUnit_mk, forall_associated, isUnit_mk, mk_mul_mk, mk_eq_mk_iff_associated, Associated.comm (x := a)] apply Iff.rfl.and @@ -1009,7 +1009,7 @@ theorem irreducible_mk {a : α} : Irreducible (Associates.mk a) ↔ Irreducible simpa using h x (y * u) (mul_assoc _ _ _) @[simp] -theorem mk_dvdNotUnit_mk_iff {a b : α} : +theorem mk_dvdNotUnit_mk_iff {a b : M} : DvdNotUnit (Associates.mk a) (Associates.mk b) ↔ DvdNotUnit a b := by simp only [DvdNotUnit, mk_ne_zero, mk_surjective.exists, isUnit_mk, mk_mul_mk, mk_eq_mk_iff_associated, Associated.comm (x := b)] @@ -1021,7 +1021,7 @@ theorem mk_dvdNotUnit_mk_iff {a b : α} : · rintro ⟨x, ⟨hx, rfl⟩⟩ use x -theorem dvdNotUnit_of_lt {a b : Associates α} (hlt : a < b) : DvdNotUnit a b := by +theorem dvdNotUnit_of_lt {a b : Associates M} (hlt : a < b) : DvdNotUnit a b := by constructor · rintro rfl apply not_lt_of_le _ hlt @@ -1033,46 +1033,46 @@ theorem dvdNotUnit_of_lt {a b : Associates α} (hlt : a < b) : DvdNotUnit a b := simp theorem irreducible_iff_prime_iff : - (∀ a : α, Irreducible a ↔ Prime a) ↔ ∀ a : Associates α, Irreducible a ↔ Prime a := by + (∀ a : M, Irreducible a ↔ Prime a) ↔ ∀ a : Associates M, Irreducible a ↔ Prime a := by simp_rw [forall_associated, irreducible_mk, prime_mk] end CommMonoidWithZero section CancelCommMonoidWithZero -variable [CancelCommMonoidWithZero α] +variable [CancelCommMonoidWithZero M] -instance instPartialOrder : PartialOrder (Associates α) where +instance instPartialOrder : PartialOrder (Associates M) where le_antisymm := mk_surjective.forall₂.2 fun _a _b hab hba => mk_eq_mk_iff_associated.2 <| associated_of_dvd_dvd (dvd_of_mk_le_mk hab) (dvd_of_mk_le_mk hba) -instance instCancelCommMonoidWithZero : CancelCommMonoidWithZero (Associates α) := - { (by infer_instance : CommMonoidWithZero (Associates α)) with +instance instCancelCommMonoidWithZero : CancelCommMonoidWithZero (Associates M) := + { (by infer_instance : CommMonoidWithZero (Associates M)) with mul_left_cancel_of_ne_zero := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ ha h rcases Quotient.exact' h with ⟨u, hu⟩ have hu : a * (b * ↑u) = a * c := by rwa [← mul_assoc] exact Quotient.sound' ⟨u, mul_left_cancel₀ (mk_ne_zero.1 ha) hu⟩ } -theorem _root_.associates_irreducible_iff_prime [DecompositionMonoid α] {p : Associates α} : +theorem _root_.associates_irreducible_iff_prime [DecompositionMonoid M] {p : Associates M} : Irreducible p ↔ Prime p := irreducible_iff_prime -instance : NoZeroDivisors (Associates α) := by infer_instance +instance : NoZeroDivisors (Associates M) := by infer_instance -theorem le_of_mul_le_mul_left (a b c : Associates α) (ha : a ≠ 0) : a * b ≤ a * c → b ≤ c +theorem le_of_mul_le_mul_left (a b c : Associates M) (ha : a ≠ 0) : a * b ≤ a * c → b ≤ c | ⟨d, hd⟩ => ⟨d, mul_left_cancel₀ ha <| by rwa [← mul_assoc]⟩ -theorem one_or_eq_of_le_of_prime {p m : Associates α} (hp : Prime p) (hle : m ≤ p) : +theorem one_or_eq_of_le_of_prime {p m : Associates M} (hp : Prime p) (hle : m ≤ p) : m = 1 ∨ m = p := by rcases mk_surjective p with ⟨p, rfl⟩ rcases mk_surjective m with ⟨m, rfl⟩ simpa [mk_eq_mk_iff_associated, Associated.comm, -Quotient.eq] using (prime_mk.1 hp).irreducible.dvd_iff.mp (mk_le_mk_iff_dvd.1 hle) -theorem dvdNotUnit_iff_lt {a b : Associates α} : DvdNotUnit a b ↔ a < b := +theorem dvdNotUnit_iff_lt {a b : Associates M} : DvdNotUnit a b ↔ a < b := dvd_and_not_dvd_iff.symm -theorem le_one_iff {p : Associates α} : p ≤ 1 ↔ p = 1 := by rw [← Associates.bot_eq_one, le_bot_iff] +theorem le_one_iff {p : Associates M} : p ≤ 1 ↔ p = 1 := by rw [← Associates.bot_eq_one, le_bot_iff] end CancelCommMonoidWithZero @@ -1080,20 +1080,20 @@ end Associates section CommMonoidWithZero -theorem DvdNotUnit.isUnit_of_irreducible_right [CommMonoidWithZero α] {p q : α} +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 α] {p q : α} (hp : ¬IsUnit p) +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 α] {p q : α} (hp : DvdNotUnit p q) : ¬IsUnit q := by +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 α] [Nontrivial α] {p q r : α} +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' obtain ⟨hp, x, hx⟩ := h @@ -1104,27 +1104,27 @@ end CommMonoidWithZero section CancelCommMonoidWithZero -theorem isUnit_of_associated_mul [CancelCommMonoidWithZero α] {p b : α} (h : Associated (p * b) p) +theorem isUnit_of_associated_mul [CancelCommMonoidWithZero M] {p b : M} (h : Associated (p * b) p) (hp : p ≠ 0) : IsUnit b := by obtain ⟨a, ha⟩ := h refine isUnit_of_mul_eq_one b a ((mul_right_inj' hp).mp ?_) rwa [← mul_assoc, mul_one] -theorem DvdNotUnit.not_associated [CancelCommMonoidWithZero α] {p q : α} (h : DvdNotUnit p q) : +theorem DvdNotUnit.not_associated [CancelCommMonoidWithZero M] {p q : M} (h : DvdNotUnit p q) : ¬Associated p q := by rintro ⟨a, rfl⟩ obtain ⟨hp, x, hx, hx'⟩ := h rcases (mul_right_inj' hp).mp hx' with rfl exact hx a.isUnit -theorem DvdNotUnit.ne [CancelCommMonoidWithZero α] {p q : α} (h : DvdNotUnit p q) : p ≠ q := by +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 α] {q : α} (hq : ¬IsUnit q) +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') @@ -1133,11 +1133,11 @@ theorem pow_injective_of_not_isUnit [CancelCommMonoidWithZero α] {q : α} (hq : @[deprecated (since := "2024-09-22")] alias pow_injective_of_not_unit := pow_injective_of_not_isUnit -theorem pow_inj_of_not_isUnit [CancelCommMonoidWithZero α] {q : α} (hq : ¬IsUnit q) +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 α] {p q : α} (hp : Prime p) (n : ℕ) : +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 | zero => From 2e75a7513e882384e3f56b0c765fff3cb58ea61e Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Mon, 30 Sep 2024 02:32:56 +0000 Subject: [PATCH 051/174] chore(Algebra/Ring): remove variables unused in section (#17178) Perhaps we should have a linter for this... @adomani? --- Mathlib/Algebra/Ring/Defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 7c35751123577..171d45485e20d 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -371,7 +371,7 @@ end NonAssocRing section Ring -variable [Ring α] {a b c d e : α} +variable [Ring α] -- A (unital, associative) ring is a not-necessarily-unital ring -- see Note [lower instance priority] From cdd209aa9627f7935ea1dffbcd0015bcb4b7722e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 02:32:57 +0000 Subject: [PATCH 052/174] =?UTF-8?q?feat:=20When=20`1=20+=20a=20=E2=89=A4?= =?UTF-8?q?=20(1=20-=20b)=20*=20(1=20+=20c)`=20(#17214)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In analysis, we often manipulate quantities near `1` as if they were additive in their difference to `1`. This PR adds lemmas supporting that kind of reasoning. From LeanAPAP --- Mathlib/Algebra/Order/Ring/Defs.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index 7e1dda2f4617c..2a1301cc48fc2 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -222,6 +222,22 @@ instance (priority := 100) OrderedRing.toOrderedSemiring : OrderedSemiring α := mul_le_mul_of_nonneg_right := fun a b c h hc => by simpa only [sub_mul, sub_nonneg] using OrderedRing.mul_nonneg _ _ (sub_nonneg.2 h) hc } +lemma one_add_le_one_sub_mul_one_add (h : a + b + b * c ≤ c) : 1 + a ≤ (1 - b) * (1 + c) := by + rw [one_sub_mul, mul_one_add, le_sub_iff_add_le, add_assoc, ← add_assoc a] + gcongr + +lemma one_add_le_one_add_mul_one_sub (h : a + c + b * c ≤ b) : 1 + a ≤ (1 + b) * (1 - c) := by + rw [mul_one_sub, one_add_mul, le_sub_iff_add_le, add_assoc, ← add_assoc a] + gcongr + +lemma one_sub_le_one_sub_mul_one_add (h : b + b * c ≤ a + c) : 1 - a ≤ (1 - b) * (1 + c) := by + rw [one_sub_mul, mul_one_add, sub_le_sub_iff, add_assoc, add_comm c] + gcongr + +lemma one_sub_le_one_add_mul_one_sub (h : c + b * c ≤ a + b) : 1 - a ≤ (1 + b) * (1 - c) := by + rw [mul_one_sub, one_add_mul, sub_le_sub_iff, add_assoc, add_comm b] + gcongr + end OrderedRing section OrderedCommRing From 07a48737b711f38253bffe0be2e236d140ed6638 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 02:32:58 +0000 Subject: [PATCH 053/174] chore(Data/ZMod/Parity): Delete (#17236) The files can be moved earlier at no cost --- Mathlib.lean | 1 - .../Combinatorics/SimpleGraph/DegreeSum.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 11 +++++++ Mathlib/Data/ZMod/Parity.lean | 33 ------------------- Mathlib/GroupTheory/Coxeter/Inversion.lean | 1 - 5 files changed, 12 insertions(+), 36 deletions(-) delete mode 100644 Mathlib/Data/ZMod/Parity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1c70c47509e0a..d12861a52ef12 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2616,7 +2616,6 @@ import Mathlib.Data.ZMod.Defs import Mathlib.Data.ZMod.Factorial import Mathlib.Data.ZMod.IntUnitsPower import Mathlib.Data.ZMod.Module -import Mathlib.Data.ZMod.Parity import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.Aliases diff --git a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean index 868f25e9ef0db..4b09938bdf173 100644 --- a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean @@ -6,7 +6,7 @@ Authors: Kyle Miller import Mathlib.Algebra.BigOperators.Ring import Mathlib.Combinatorics.SimpleGraph.Dart import Mathlib.Combinatorics.SimpleGraph.Finite -import Mathlib.Data.ZMod.Parity +import Mathlib.Data.ZMod.Basic /-! # Degree-sum formula and handshaking lemma diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 587f8afc8aa50..8c9d56fcedfec 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -826,6 +826,17 @@ theorem eq_iff_modEq_nat (n : ℕ) {a b : ℕ} : (a : ZMod n) = b ↔ a ≡ b [M · rw [Fin.ext_iff, Nat.ModEq, ← val_natCast, ← val_natCast] exact Iff.rfl +theorem eq_zero_iff_even {n : ℕ} : (n : ZMod 2) = 0 ↔ Even n := + (CharP.cast_eq_zero_iff (ZMod 2) 2 n).trans even_iff_two_dvd.symm + +theorem eq_one_iff_odd {n : ℕ} : (n : ZMod 2) = 1 ↔ Odd n := by + rw [← @Nat.cast_one (ZMod 2), ZMod.eq_iff_modEq_nat, Nat.odd_iff, Nat.ModEq] + +theorem ne_zero_iff_odd {n : ℕ} : (n : ZMod 2) ≠ 0 ↔ Odd n := by + constructor <;> + · contrapose + simp [eq_zero_iff_even] + theorem coe_mul_inv_eq_one {n : ℕ} (x : ℕ) (h : Nat.Coprime x n) : ((x : ZMod n) * (x : ZMod n)⁻¹) = 1 := by rw [Nat.Coprime, Nat.gcd_comm, Nat.gcd_rec] at h diff --git a/Mathlib/Data/ZMod/Parity.lean b/Mathlib/Data/ZMod/Parity.lean deleted file mode 100644 index b6686f1e23ec5..0000000000000 --- a/Mathlib/Data/ZMod/Parity.lean +++ /dev/null @@ -1,33 +0,0 @@ -/- -Copyright (c) 2020 Kyle Miller. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kyle Miller --/ -import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Data.ZMod.Basic - -/-! -# Relating parity to natural numbers mod 2 - -This module provides lemmas relating `ZMod 2` to `Even` and `Odd`. - -## Tags - -parity, zmod, even, odd --/ - - -namespace ZMod - -theorem eq_zero_iff_even {n : ℕ} : (n : ZMod 2) = 0 ↔ Even n := - (CharP.cast_eq_zero_iff (ZMod 2) 2 n).trans even_iff_two_dvd.symm - -theorem eq_one_iff_odd {n : ℕ} : (n : ZMod 2) = 1 ↔ Odd n := by - rw [← @Nat.cast_one (ZMod 2), ZMod.eq_iff_modEq_nat, Nat.odd_iff, Nat.ModEq] - -theorem ne_zero_iff_odd {n : ℕ} : (n : ZMod 2) ≠ 0 ↔ Odd n := by - constructor <;> - · contrapose - simp [eq_zero_iff_even] - -end ZMod diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index b07693855316a..18946d5bd27dc 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mitchell Lee -/ import Mathlib.GroupTheory.Coxeter.Length -import Mathlib.Data.ZMod.Parity import Mathlib.Data.List.GetD /-! From eed3300c27c9f168d53e13bb198a92a147b671d0 Mon Sep 17 00:00:00 2001 From: FR Date: Mon, 30 Sep 2024 03:28:53 +0000 Subject: [PATCH 054/174] doc(Data/Rat/Cast/Defs): update docs (#17268) The notation has been moved to `Batteries.Data.Rat`. --- Mathlib/Data/Rat/Cast/Defs.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index 643f4426c88c1..739dbc17cb0ce 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -19,10 +19,6 @@ import Mathlib.Data.Rat.Lemmas We define the canonical injection from ℚ into an arbitrary division ring and prove various casting lemmas showing the well-behavedness of this injection. -## Notations - -- `/.` is infix notation for `Rat.divInt`. - ## Tags rat, rationals, field, ℚ, numerator, denominator, num, denom, cast, coercion, casting From 39b61a6aaa6327f75e5d27b5320e2fc4856303ca Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 30 Sep 2024 08:52:11 +0000 Subject: [PATCH 055/174] chore(Data/Matrix): some trivial lemmas about `row` and `col` (#17243) Co-authored-by: Eric Wieser --- Mathlib/Data/Matrix/RowCol.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Data/Matrix/RowCol.lean b/Mathlib/Data/Matrix/RowCol.lean index 2529d88fd19ee..6ce7feb0bcaa5 100644 --- a/Mathlib/Data/Matrix/RowCol.lean +++ b/Mathlib/Data/Matrix/RowCol.lean @@ -136,6 +136,16 @@ theorem row_mulVec [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α ext rfl +theorem row_mulVec_eq_const [Fintype m] [NonUnitalNonAssocSemiring α] (v w : m → α) : + Matrix.row ι v *ᵥ w = Function.const _ (v ⬝ᵥ w) := rfl + +theorem mulVec_col_eq_const [Fintype m] [NonUnitalNonAssocSemiring α] (v w : m → α) : + v ᵥ* Matrix.col ι w = Function.const _ (v ⬝ᵥ w) := rfl + +theorem row_mul_col [Fintype m] [Mul α] [AddCommMonoid α] (v w : m → α) : + row ι v * col ι w = of fun _ _ => v ⬝ᵥ w := + rfl + @[simp] theorem row_mul_col_apply [Fintype m] [Mul α] [AddCommMonoid α] (v w : m → α) (i j) : (row ι v * col ι w) i j = v ⬝ᵥ w := From e336e27e0824583a3919ded9a6cb5870b198bbbb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 09:29:24 +0000 Subject: [PATCH 056/174] chore(ZMod): Move module and algebra instances under `Algebra` (#17237) Those files have nothing to do under `Data` --- Mathlib.lean | 4 ++-- .../{Data/ZMod/Algebra.lean => Algebra/Algebra/ZMod.lean} | 2 +- Mathlib/Algebra/Module/Torsion.lean | 6 +++--- Mathlib/{Data/ZMod/Module.lean => Algebra/Module/ZMod.lean} | 1 - Mathlib/FieldTheory/Finite/GaloisField.lean | 2 +- Mathlib/FieldTheory/IsAlgClosed/Classification.lean | 4 ++-- Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean | 2 +- Mathlib/RingTheory/WittVector/Frobenius.lean | 4 ++-- 8 files changed, 12 insertions(+), 13 deletions(-) rename Mathlib/{Data/ZMod/Algebra.lean => Algebra/Algebra/ZMod.lean} (100%) rename Mathlib/{Data/ZMod/Module.lean => Algebra/Module/ZMod.lean} (98%) diff --git a/Mathlib.lean b/Mathlib.lean index d12861a52ef12..66017fa31543d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -32,6 +32,7 @@ import Mathlib.Algebra.Algebra.Subalgebra.Tower import Mathlib.Algebra.Algebra.Subalgebra.Unitization import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.Algebra.Unitization +import Mathlib.Algebra.Algebra.ZMod import Mathlib.Algebra.AlgebraicCard import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.Associated.OrderedCommMonoid @@ -507,6 +508,7 @@ import Mathlib.Algebra.Module.Torsion import Mathlib.Algebra.Module.ULift import Mathlib.Algebra.Module.ZLattice.Basic import Mathlib.Algebra.Module.ZLattice.Covolume +import Mathlib.Algebra.Module.ZMod import Mathlib.Algebra.MonoidAlgebra.Basic import Mathlib.Algebra.MonoidAlgebra.Defs import Mathlib.Algebra.MonoidAlgebra.Degree @@ -2609,13 +2611,11 @@ import Mathlib.Data.Vector3 import Mathlib.Data.W.Basic import Mathlib.Data.W.Cardinal import Mathlib.Data.W.Constructions -import Mathlib.Data.ZMod.Algebra import Mathlib.Data.ZMod.Basic import Mathlib.Data.ZMod.Coprime import Mathlib.Data.ZMod.Defs import Mathlib.Data.ZMod.Factorial import Mathlib.Data.ZMod.IntUnitsPower -import Mathlib.Data.ZMod.Module import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.Aliases diff --git a/Mathlib/Data/ZMod/Algebra.lean b/Mathlib/Algebra/Algebra/ZMod.lean similarity index 100% rename from Mathlib/Data/ZMod/Algebra.lean rename to Mathlib/Algebra/Algebra/ZMod.lean index 2f5737b396fd4..a04244c46ce31 100644 --- a/Mathlib/Data/ZMod/Algebra.lean +++ b/Mathlib/Algebra/Algebra/ZMod.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Data.ZMod.Basic import Mathlib.Algebra.Algebra.Defs +import Mathlib.Data.ZMod.Basic /-! # The `ZMod n`-algebra structure on rings whose characteristic divides `n` diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index ecd92053f2c94..5be854af327cb 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -3,11 +3,11 @@ Copyright (c) 2022 Pierre-Alexandre Bazin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Pierre-Alexandre Bazin -/ -import Mathlib.LinearAlgebra.Isomorphisms +import Mathlib.Algebra.DirectSum.Module +import Mathlib.Algebra.Module.ZMod import Mathlib.GroupTheory.Torsion +import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Coprime.Ideal -import Mathlib.Data.ZMod.Module -import Mathlib.Algebra.DirectSum.Module /-! # Torsion submodules diff --git a/Mathlib/Data/ZMod/Module.lean b/Mathlib/Algebra/Module/ZMod.lean similarity index 98% rename from Mathlib/Data/ZMod/Module.lean rename to Mathlib/Algebra/Module/ZMod.lean index c99250ffefff5..961af9fe2c048 100644 --- a/Mathlib/Data/ZMod/Module.lean +++ b/Mathlib/Algebra/Module/ZMod.lean @@ -5,7 +5,6 @@ Authors: Lawrence Wu -/ import Mathlib.Algebra.Module.Submodule.Lattice import Mathlib.Data.ZMod.Basic -import Mathlib.Order.OmegaCompletePartialOrder /-! # The `ZMod n`-module structure on Abelian groups whose elements have order dividing `n` diff --git a/Mathlib/FieldTheory/Finite/GaloisField.lean b/Mathlib/FieldTheory/Finite/GaloisField.lean index d26ad82dcaec4..e03d1331f1020 100644 --- a/Mathlib/FieldTheory/Finite/GaloisField.lean +++ b/Mathlib/FieldTheory/Finite/GaloisField.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Alex J. Best, Johan Commelin, Eric Rodriguez, Ruben Van de Velde -/ +import Mathlib.Algebra.Algebra.ZMod import Mathlib.Algebra.CharP.Algebra -import Mathlib.Data.ZMod.Algebra import Mathlib.FieldTheory.Finite.Basic import Mathlib.FieldTheory.Galois.Basic import Mathlib.FieldTheory.SplittingField.IsSplittingField diff --git a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean index eadc63373a37b..2f13ff3833e58 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean @@ -3,9 +3,9 @@ Copyright (c) 2022 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Polynomial.Cardinal +import Mathlib.Algebra.Algebra.ZMod import Mathlib.Algebra.MvPolynomial.Cardinal -import Mathlib.Data.ZMod.Algebra +import Mathlib.Algebra.Polynomial.Cardinal import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.RingTheory.AlgebraicIndependent diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean index 5f2544087a31a..d66fdc3f66ab9 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import Mathlib.Algebra.Algebra.ZMod import Mathlib.RingTheory.Polynomial.Cyclotomic.Roots -import Mathlib.Data.ZMod.Algebra /-! # Cyclotomic polynomials and `expand`. diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index 0492664ca9471..0d989e4ecdebe 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -3,11 +3,11 @@ 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.Algebra.ZMod import Mathlib.Data.Nat.Multiplicity -import Mathlib.Data.ZMod.Algebra +import Mathlib.FieldTheory.Perfect import Mathlib.RingTheory.WittVector.Basic import Mathlib.RingTheory.WittVector.IsPoly -import Mathlib.FieldTheory.Perfect /-! ## The Frobenius operator From 04f66c4a08fb7402b3a367217273611eecd28199 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 10:30:18 +0000 Subject: [PATCH 057/174] chore: `Finsupp.coe_finset_sum` simp (#17228) From LeanCamCombi --- Mathlib/Algebra/BigOperators/Finsupp.lean | 4 ++-- Mathlib/Data/Nat/Factorization/Defs.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index 35f43a5d0baa1..10fdcfb6979ea 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -249,12 +249,12 @@ theorem sum_apply [Zero M] [AddCommMonoid N] {f : α →₀ M} {g : α → M → finset_sum_apply _ _ _ -- Porting note: inserted ⇑ on the rhs -theorem coe_finset_sum [AddCommMonoid N] (S : Finset ι) (f : ι → α →₀ N) : +@[simp, norm_cast] theorem coe_finset_sum [AddCommMonoid N] (S : Finset ι) (f : ι → α →₀ N) : ⇑(∑ i ∈ S, f i) = ∑ i ∈ S, ⇑(f i) := map_sum (coeFnAddHom : (α →₀ N) →+ _) _ _ -- Porting note: inserted ⇑ on the rhs -theorem coe_sum [Zero M] [AddCommMonoid N] (f : α →₀ M) (g : α → M → β →₀ N) : +@[simp, norm_cast] theorem coe_sum [Zero M] [AddCommMonoid N] (f : α →₀ M) (g : α → M → β →₀ N) : ⇑(f.sum g) = f.sum fun a₁ b => ⇑(g a₁ b) := coe_finset_sum _ _ diff --git a/Mathlib/Data/Nat/Factorization/Defs.lean b/Mathlib/Data/Nat/Factorization/Defs.lean index fbabc1edc00fd..4ff0ac2cf90f5 100644 --- a/Mathlib/Data/Nat/Factorization/Defs.lean +++ b/Mathlib/Data/Nat/Factorization/Defs.lean @@ -172,7 +172,7 @@ theorem factorization_prod {α : Type*} {S : Finset α} {g : α → ℕ} (hS : · simp · intro x T hxS hTS hxT IH have hT : T.prod g ≠ 0 := prod_ne_zero_iff.mpr fun x hx => hS x (hTS hx) - simp [prod_insert hxT, sum_insert hxT, ← IH, factorization_mul (hS x hxS) hT] + simp [prod_insert hxT, sum_insert hxT, IH, factorization_mul (hS x hxS) hT] /-- For any `p`, the power of `p` in `n^k` is `k` times the power in `n` -/ @[simp] From c07cbe11e5a8efdd2ae729499e59ac1eaf0c593b Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 30 Sep 2024 11:16:56 +0000 Subject: [PATCH 058/174] chore(Analysis/CStarAlgebra/Spectrum): generalize type class assumptions (#17267) --- Mathlib/Analysis/CStarAlgebra/Spectrum.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean index c34a1be9cfe71..337a423b3d120 100644 --- a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -263,9 +263,9 @@ end NonUnitalStarAlgHom namespace StarAlgEquiv -variable {F A B : Type*} [NormedRing A] [NormedSpace ℂ A] [SMulCommClass ℂ A A] +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 [NormedRing B] [NormedSpace ℂ B] [SMulCommClass ℂ B B] [IsScalarTower ℂ B B] +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 [NonUnitalAlgEquivClass F ℂ A B] [StarHomClass F A B] From 71a892158a1a621d92577030074e52d7dff0806a Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 30 Sep 2024 11:16:57 +0000 Subject: [PATCH 059/174] chore: remove shadowed variable name (#17272) The variable name `H` was repeated. The ambiguity could potentially be an issue for using `apply lemma_name (H := x)`. Found while working on the unused variable command linter. --- Mathlib/RingTheory/IsTensorProduct.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index 942b0ab45ddf4..d146835dbbfcd 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -424,7 +424,7 @@ noncomputable def Algebra.pushoutDesc [H : Algebra.IsPushout R S R' S'] {A : Typ rw [mul_add, map_add, map_add, mul_add, e₁, e₂] @[simp] -theorem Algebra.pushoutDesc_left [H : Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] +theorem Algebra.pushoutDesc_left [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] [Algebra R A] (f : S →ₐ[R] A) (g : R' →ₐ[R] A) (H) (x : S) : Algebra.pushoutDesc S' f g H (algebraMap S S' x) = f x := by letI := Module.compHom A f.toRingHom @@ -442,7 +442,7 @@ theorem Algebra.lift_algHom_comp_left [Algebra.IsPushout R S R' S'] {A : Type*} AlgHom.ext fun x => (Algebra.pushoutDesc_left S' f g H x : _) @[simp] -theorem Algebra.pushoutDesc_right [H : Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] +theorem Algebra.pushoutDesc_right [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] [Algebra R A] (f : S →ₐ[R] A) (g : R' →ₐ[R] A) (H) (x : R') : Algebra.pushoutDesc S' f g H (algebraMap R' S' x) = g x := letI := Module.compHom A f.toRingHom From dfc92350bbbd48a7af8224db543dfcd4eba0c4b1 Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Mon, 30 Sep 2024 12:08:26 +0000 Subject: [PATCH 060/174] chore(Tactic/Linter): do not import `TextBased` linter (#17187) `Mathlib.Tactic.Linter.TextBased` is only used by the script `lake exe lint-style` and therefore does not to be imported by Mathlib. --- Mathlib/Tactic/Linter.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Tactic/Linter.lean b/Mathlib/Tactic/Linter.lean index e40a62f821d5a..e5c4da4423d3c 100644 --- a/Mathlib/Tactic/Linter.lean +++ b/Mathlib/Tactic/Linter.lean @@ -11,4 +11,3 @@ This file is ignored by `shake`: import Mathlib.Tactic.Linter.FlexibleLinter import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.MinImports -import Mathlib.Tactic.Linter.TextBased From 13d70fadd35fe56699f6d38a88fa541085158669 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 30 Sep 2024 13:03:36 +0000 Subject: [PATCH 061/174] chore(Tactic/GeneraliseProofs): reduce Inhabited to Nonempty (#17125) This seems to work, but I am not sure if this change is fully fine, or there are side effects I don't know about. --- Mathlib/Tactic/GeneralizeProofs.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Tactic/GeneralizeProofs.lean b/Mathlib/Tactic/GeneralizeProofs.lean index 7acd8bef29e3e..645637ae46a0f 100644 --- a/Mathlib/Tactic/GeneralizeProofs.lean +++ b/Mathlib/Tactic/GeneralizeProofs.lean @@ -341,7 +341,7 @@ This continuation `k` is passed The `propToFVar` map is updated with the new proposition fvars. -/ -partial def withGeneralizedProofs {α : Type} [Inhabited α] (e : Expr) (ty? : Option Expr) +partial def withGeneralizedProofs {α : Type} [Nonempty α] (e : Expr) (ty? : Option Expr) (k : Array Expr → Array Expr → Expr → MGen α) : MGen α := do let propToFVar := (← get).propToFVar @@ -351,7 +351,7 @@ partial def withGeneralizedProofs {α : Type} [Inhabited α] (e : Expr) (ty? : O post-abstracted{indentD e}\nnew generalizations: {generalizations}" let rec /-- Core loop for `withGeneralizedProofs`, adds generalizations one at a time. -/ - go [Inhabited α] (i : Nat) (fvars pfs : Array Expr) + go [Nonempty α] (i : Nat) (fvars pfs : Array Expr) (proofToFVar propToFVar : ExprMap Expr) : MGen α := do if h : i < generalizations.size then let (ty, pf) := generalizations[i] From d86c16309c2287d6899defc123c50c0cd809081a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 30 Sep 2024 13:03:37 +0000 Subject: [PATCH 062/174] feat: add missing List.foldr lemmas (#17222) --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Data/List/Basic.lean | 6 +++++ Mathlib/Data/List/Perm.lean | 26 +++++++++++++++++++- 3 files changed, 32 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index f097934c26635..59be51355e078 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -342,7 +342,7 @@ lemma prod_map_erase [DecidableEq α] (f : α → M) {a} : · simp only [map, erase_cons_tail (not_beq_of_ne ne.symm), prod_cons, prod_map_erase _ h, mul_left_comm (f a) (f b)] -@[to_additive] lemma Perm.prod_eq (h : Perm l₁ l₂) : prod l₁ = prod l₂ := h.fold_op_eq +@[to_additive] lemma Perm.prod_eq (h : Perm l₁ l₂) : prod l₁ = prod l₂ := h.foldl_op_eq @[to_additive] lemma prod_reverse (l : List M) : prod l.reverse = prod l := (reverse_perm l).prod_eq diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index c239c183abaad..b4562695ae718 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1373,6 +1373,12 @@ theorem foldl_op_eq_op_foldr_assoc : | a :: l, a₁, a₂ => by simp only [foldl_cons, foldr_cons, foldl_assoc, ha.assoc]; rw [foldl_op_eq_op_foldr_assoc] +theorem foldr_assoc : ∀ {l : List α} {a₁ a₂}, l.foldr op (op a₁ a₂) = op (l.foldr op a₁) a₂ + | [], a₁, a₂ => rfl + | a :: l, a₁, a₂ => by + simp only [foldr_cons, ha.assoc] + rw [foldr_assoc] + variable [hc : Std.Commutative op] theorem foldl_assoc_comm_cons {l : List α} {a₁ a₂} : ((a₁ :: l) <*> a₂) = a₁ ⋆ l <*> a₂ := by diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 36d55e987eaee..fdff37f5163e5 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -51,6 +51,25 @@ theorem Perm.subset_congr_left {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l theorem Perm.subset_congr_right {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l₃ ⊆ l₁ ↔ l₃ ⊆ l₂ := ⟨fun h' => h'.trans h.subset, fun h' => h'.trans h.symm.subset⟩ +/-- Variant of `Perm.foldr_eq` with explicit commutativity argument. -/ +theorem Perm.foldr_eq' {f : α → β → β} {l₁ l₂ : List α} (p : l₁ ~ l₂) + (comm : ∀ x ∈ l₁, ∀ y ∈ l₁, ∀ z, f y (f x z) = f x (f y z)) + (init : β) : foldr f init l₁ = foldr f init l₂ := by + induction p using recOnSwap' generalizing init with + | nil => simp + | cons x _p IH => + simp only [foldr] + congr 1 + apply IH; intros; apply comm <;> exact .tail _ ‹_› + | swap' x y _p IH => + simp only [foldr] + rw [comm x (.tail _ <| .head _) y (.head _)] + congr 2 + apply IH; intros; apply comm <;> exact .tail _ (.tail _ ‹_›) + | trans p₁ _p₂ IH₁ IH₂ => + refine (IH₁ comm init).trans (IH₂ ?_ _) + intros; apply comm <;> apply p₁.symm.subset <;> assumption + section Rel open Relator @@ -166,9 +185,14 @@ local notation a " * " b => op a b local notation l " <*> " a => foldl op a l -theorem Perm.fold_op_eq {l₁ l₂ : List α} {a : α} (h : l₁ ~ l₂) : (l₁ <*> a) = l₂ <*> a := +theorem Perm.foldl_op_eq {l₁ l₂ : List α} {a : α} (h : l₁ ~ l₂) : (l₁ <*> a) = l₂ <*> a := h.foldl_eq _ +theorem Perm.foldr_op_eq {l₁ l₂ : List α} {a : α} (h : l₁ ~ l₂) : l₁.foldr op a = l₂.foldr op a := + h.foldr_eq _ + +@[deprecated (since := "2024-09-28")] alias Perm.fold_op_eq := Perm.foldl_op_eq + end theorem perm_option_to_list {o₁ o₂ : Option α} : o₁.toList ~ o₂.toList ↔ o₁ = o₂ := by From 2db234e2d5db1c1df0efbbaf581cc1c46194a12c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 14:06:03 +0000 Subject: [PATCH 063/174] chore: Rename `DenseEmbedding` to `IsDenseEmbedding` (#17247) `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 `DenseEmbedding` 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. --- .../Analysis/InnerProductSpace/OfNorm.lean | 2 +- .../Function/SimpleFuncDenseLp.lean | 13 +++-- Mathlib/NumberTheory/Liouville/Residual.lean | 2 +- Mathlib/Topology/Algebra/UniformField.lean | 2 +- .../Topology/Compactification/OnePoint.lean | 7 ++- Mathlib/Topology/DenseEmbedding.lean | 51 +++++++++++-------- Mathlib/Topology/Homeomorph.lean | 10 +++- Mathlib/Topology/Instances/Complex.lean | 4 +- Mathlib/Topology/Instances/Rat.lean | 9 ++-- Mathlib/Topology/Instances/RatLemmas.lean | 2 +- .../Topology/Instances/RealVectorSpace.lean | 2 +- Mathlib/Topology/StoneCech.lean | 5 +- .../Topology/UniformSpace/CompareReals.lean | 2 +- Mathlib/Topology/UniformSpace/Completion.lean | 14 +++-- .../UniformSpace/UniformEmbedding.lean | 20 +++++--- scripts/no_lints_prime_decls.txt | 2 +- 16 files changed, 92 insertions(+), 55 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index 2616c07526d0b..48732656917c7 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -269,7 +269,7 @@ private theorem real_prop (r : ℝ) : innerProp' E (r : 𝕜) := by intro x y revert r rw [← Function.funext_iff] - refine Rat.denseEmbedding_coe_real.dense.equalizer ?_ ?_ (funext fun X => ?_) + 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 · simp only [Function.comp_apply, RCLike.ofReal_ratCast, rat_prop _ _] diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index d2aa0a504c720..29f9a7d7b2d9b 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -23,7 +23,7 @@ by a sequence of simple functions. measurable and `Memℒp` (for `p < ∞`), then the simple functions `SimpleFunc.approxOn f hf s 0 h₀ n` may be considered as elements of `Lp E p μ`, and they tend in Lᵖ to `f`. -* `Lp.simpleFunc.denseEmbedding`: the embedding `coeToLp` of the `Lp` simple functions into +* `Lp.simpleFunc.isDenseEmbedding`: the embedding `coeToLp` of the `Lp` simple functions into `Lp` is dense. * `Lp.simpleFunc.induction`, `Lp.induction`, `Memℒp.induction`, `Integrable.induction`: to prove a predicate for all elements of one of these classes of functions, it suffices to check that it @@ -685,10 +685,10 @@ protected theorem uniformEmbedding : UniformEmbedding ((↑) : Lp.simpleFunc E p protected theorem uniformInducing : UniformInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := simpleFunc.uniformEmbedding.toUniformInducing -protected theorem denseEmbedding (hp_ne_top : p ≠ ∞) : - DenseEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := by +lemma isDenseEmbedding (hp_ne_top : p ≠ ∞) : + IsDenseEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := by borelize E - apply simpleFunc.uniformEmbedding.denseEmbedding + apply simpleFunc.uniformEmbedding.isDenseEmbedding intro f rw [mem_closure_iff_seq_limit] have hfi' : Memℒp f p μ := Lp.memℒp f @@ -703,9 +703,12 @@ protected theorem denseEmbedding (hp_ne_top : p ≠ ∞) : convert SimpleFunc.tendsto_approxOn_range_Lp hp_ne_top (Lp.stronglyMeasurable f).measurable hfi' rw [toLp_coeFn f (Lp.memℒp f)] +@[deprecated (since := "2024-09-30")] +alias denseEmbedding := isDenseEmbedding + protected theorem isDenseInducing (hp_ne_top : p ≠ ∞) : IsDenseInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - (simpleFunc.denseEmbedding hp_ne_top).toIsDenseInducing + (simpleFunc.isDenseEmbedding hp_ne_top).toIsDenseInducing protected theorem denseRange (hp_ne_top : p ≠ ∞) : DenseRange ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := diff --git a/Mathlib/NumberTheory/Liouville/Residual.lean b/Mathlib/NumberTheory/Liouville/Residual.lean index 4d72a848d850a..38feb5335e927 100644 --- a/Mathlib/NumberTheory/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Liouville/Residual.lean @@ -53,7 +53,7 @@ theorem setOf_liouville_eq_irrational_inter_iInter_iUnion : theorem eventually_residual_liouville : ∀ᶠ x in residual ℝ, Liouville x := by rw [Filter.Eventually, setOf_liouville_eq_irrational_inter_iInter_iUnion] refine eventually_residual_irrational.and ?_ - refine residual_of_dense_Gδ ?_ (Rat.denseEmbedding_coe_real.dense.mono ?_) + refine residual_of_dense_Gδ ?_ (Rat.isDenseEmbedding_coe_real.dense.mono ?_) · exact .iInter fun n => IsOpen.isGδ <| isOpen_iUnion fun a => isOpen_iUnion fun b => isOpen_iUnion fun _hb => isOpen_ball · rintro _ ⟨r, rfl⟩ diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 59733a4968e23..27cc2841c5eb2 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -111,7 +111,7 @@ theorem coe_inv (x : K) : (x : hat K)⁻¹ = ((x⁻¹ : K) : hat K) := by · conv_lhs => dsimp [Inv.inv] rw [if_neg] · exact hatInv_extends h - · exact fun H => h (denseEmbedding_coe.inj H) + · exact fun H => h (isDenseEmbedding_coe.inj H) variable [UniformAddGroup K] diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index cb7d3c40f534c..ffe89b1a9a2a8 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -426,9 +426,12 @@ theorem denseRange_coe [NoncompactSpace X] : DenseRange ((↑) : X → OnePoint rw [DenseRange, ← compl_infty] exact dense_compl_singleton _ -theorem denseEmbedding_coe [NoncompactSpace X] : DenseEmbedding ((↑) : X → OnePoint X) := +theorem isDenseEmbedding_coe [NoncompactSpace X] : IsDenseEmbedding ((↑) : X → OnePoint X) := { openEmbedding_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 @@ -507,7 +510,7 @@ example [WeaklyLocallyCompactSpace X] [T2Space X] : T4Space (OnePoint X) := infe /-- If `X` is not a compact space, then `OnePoint X` is a connected space. -/ instance [PreconnectedSpace X] [NoncompactSpace X] : ConnectedSpace (OnePoint X) where - toPreconnectedSpace := denseEmbedding_coe.toIsDenseInducing.preconnectedSpace + toPreconnectedSpace := isDenseEmbedding_coe.toIsDenseInducing.preconnectedSpace toNonempty := inferInstance /-- If `X` is an infinite type with discrete topology (e.g., `ℕ`), then the identity map from diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 78c6173ce6866..84bfa9ec96c78 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -11,9 +11,9 @@ import Mathlib.Topology.Bases This file defines three properties of functions: -* `DenseRange f` means `f` has dense image; -* `IsDenseInducing i` means `i` is also `Inducing`, namely it induces the topology on its codomain; -* `DenseEmbedding e` means `e` is further an `Embedding`, namely it is injective and `Inducing`. +* `DenseRange f` means `f` has dense image; +* `IsDenseInducing i` means `i` is also `Inducing`, namely it induces the topology on its codomain; +* `IsDenseEmbedding e` means `e` is further an `Embedding`, namely it is injective and `Inducing`. The main theorem `continuous_extend` gives a criterion for a function `f : X → Z` to a T₃ space Z to extend along a dense embedding @@ -202,37 +202,40 @@ theorem mk' (i : α → β) (c : Continuous i) (dense : ∀ x, x ∈ closure (ra end IsDenseInducing /-- A dense embedding is an embedding with dense image. -/ -structure DenseEmbedding [TopologicalSpace α] [TopologicalSpace β] (e : α → β) extends +structure IsDenseEmbedding [TopologicalSpace α] [TopologicalSpace β] (e : α → β) extends IsDenseInducing e : Prop where /-- A dense embedding is injective. -/ inj : Function.Injective e -theorem DenseEmbedding.mk' [TopologicalSpace α] [TopologicalSpace β] (e : α → β) (c : Continuous e) +lemma IsDenseEmbedding.mk' [TopologicalSpace α] [TopologicalSpace β] (e : α → β) (c : Continuous e) (dense : DenseRange e) (inj : Function.Injective e) - (H : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 (e a), ∀ b, e b ∈ t → b ∈ s) : DenseEmbedding e := + (H : ∀ (a : α), ∀ s ∈ 𝓝 a, ∃ t ∈ 𝓝 (e a), ∀ b, e b ∈ t → b ∈ s) : IsDenseEmbedding e := { IsDenseInducing.mk' e c dense H with inj } -namespace DenseEmbedding +@[deprecated (since := "2024-09-30")] +alias DenseEmbedding.mk' := IsDenseEmbedding.mk' + +namespace IsDenseEmbedding open TopologicalSpace variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] variable {e : α → β} -theorem inj_iff (de : DenseEmbedding e) {x y} : e x = e y ↔ x = y := +theorem inj_iff (de : IsDenseEmbedding e) {x y} : e x = e y ↔ x = y := de.inj.eq_iff -theorem to_embedding (de : DenseEmbedding e) : Embedding e := +theorem to_embedding (de : IsDenseEmbedding e) : Embedding e := { induced := de.induced inj := de.inj } -/-- If the domain of a `DenseEmbedding` is a separable space, then so is its codomain. -/ -protected theorem separableSpace [SeparableSpace α] (de : DenseEmbedding e) : SeparableSpace β := +/-- If the domain of a `IsDenseEmbedding` is a separable space, then so is its codomain. -/ +protected theorem separableSpace [SeparableSpace α] (de : IsDenseEmbedding e) : SeparableSpace β := de.toIsDenseInducing.separableSpace /-- The product of two dense embeddings is a dense embedding. -/ -protected theorem prod {e₁ : α → β} {e₂ : γ → δ} (de₁ : DenseEmbedding e₁) - (de₂ : DenseEmbedding e₂) : DenseEmbedding fun p : α × γ => (e₁ p.1, e₂ p.2) := +protected theorem prod {e₁ : α → β} {e₂ : γ → δ} (de₁ : IsDenseEmbedding e₁) + (de₂ : IsDenseEmbedding e₂) : IsDenseEmbedding fun p : α × γ => (e₁ p.1, e₂ p.2) := { de₁.toIsDenseInducing.prod de₂.toIsDenseInducing with inj := de₁.inj.prodMap de₂.inj } @@ -242,8 +245,8 @@ def subtypeEmb {α : Type*} (p : α → Prop) (e : α → β) (x : { x // p x }) { x // x ∈ closure (e '' { x | p x }) } := ⟨e x, subset_closure <| mem_image_of_mem e x.prop⟩ -protected theorem subtype (de : DenseEmbedding e) (p : α → Prop) : - DenseEmbedding (subtypeEmb p e) := +protected theorem subtype (de : IsDenseEmbedding e) (p : α → Prop) : + IsDenseEmbedding (subtypeEmb p e) := { dense := dense_iff_closure_eq.2 <| by ext ⟨x, hx⟩ @@ -255,18 +258,24 @@ protected theorem subtype (de : DenseEmbedding e) (p : α → Prop) : simp [subtypeEmb, nhds_subtype_eq_comap, de.toInducing.nhds_eq_comap, comap_comap, Function.comp_def] } -theorem dense_image (de : DenseEmbedding e) {s : Set α} : Dense (e '' s) ↔ Dense s := +theorem dense_image (de : IsDenseEmbedding e) {s : Set α} : Dense (e '' s) ↔ Dense s := de.toIsDenseInducing.dense_image -end DenseEmbedding - -theorem denseEmbedding_id {α : Type*} [TopologicalSpace α] : DenseEmbedding (id : α → α) := +protected lemma id {α : Type*} [TopologicalSpace α] : IsDenseEmbedding (id : α → α) := { embedding_id with dense := denseRange_id } -theorem Dense.denseEmbedding_val [TopologicalSpace α] {s : Set α} (hs : Dense s) : - DenseEmbedding ((↑) : s → α) := +end IsDenseEmbedding + +@[deprecated (since := "2024-09-30")] +alias denseEmbedding_id := IsDenseEmbedding.id + +theorem Dense.isDenseEmbedding_val [TopologicalSpace α] {s : Set α} (hs : Dense s) : + IsDenseEmbedding ((↑) : s → α) := { embedding_subtype_val with dense := hs.denseRange_val } +@[deprecated (since := "2024-09-30")] +alias Dense.denseEmbedding_val := Dense.isDenseEmbedding_val + theorem isClosed_property [TopologicalSpace β] {e : α → β} {p : β → Prop} (he : DenseRange e) (hp : IsClosed { x | p x }) (h : ∀ a, p (e a)) : ∀ b, p b := have : univ ⊆ { b | p b } := diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index bc14d3525fda0..228d23d3112d8 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -308,9 +308,12 @@ protected theorem t2Space [T2Space X] (h : X ≃ₜ Y) : T2Space Y := protected theorem t3Space [T3Space X] (h : X ≃ₜ Y) : T3Space Y := h.symm.embedding.t3Space -protected theorem denseEmbedding (h : X ≃ₜ Y) : DenseEmbedding h := +theorem isDenseEmbedding (h : X ≃ₜ Y) : IsDenseEmbedding h := { h.embedding with dense := h.surjective.denseRange } +@[deprecated (since := "2024-09-30")] +alias denseEmbedding := isDenseEmbedding + @[simp] theorem isOpen_preimage (h : X ≃ₜ Y) {s : Set Y} : IsOpen (h ⁻¹' s) ↔ IsOpen s := h.quotientMap.isOpen_preimage @@ -901,7 +904,10 @@ 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 protected lemma closedEmbedding : ClosedEmbedding f := (hf.homeomorph f).closedEmbedding -protected lemma denseEmbedding : DenseEmbedding f := (hf.homeomorph f).denseEmbedding +lemma isDenseEmbedding : IsDenseEmbedding f := (hf.homeomorph f).isDenseEmbedding + +@[deprecated (since := "2024-09-30")] +alias denseEmbedding := isDenseEmbedding end IsHomeomorph diff --git a/Mathlib/Topology/Instances/Complex.lean b/Mathlib/Topology/Instances/Complex.lean index 8cfe0e613fcb3..6b8e3c0862a14 100644 --- a/Mathlib/Topology/Instances/Complex.lean +++ b/Mathlib/Topology/Instances/Complex.lean @@ -39,7 +39,7 @@ theorem Complex.subfield_eq_of_closed {K : Subfield ℂ} (hc : IsClosed (K : Set simp only [Function.comp_apply, ofReal_ratCast, SetLike.mem_coe, SubfieldClass.ratCast_mem] nth_rw 1 [range_comp] refine subset_trans ?_ (image_closure_subset_closure_image continuous_ofReal) - rw [DenseRange.closure_range Rat.denseEmbedding_coe_real.dense] + rw [DenseRange.closure_range Rat.isDenseEmbedding_coe_real.dense] simp only [image_univ] rfl @@ -106,7 +106,7 @@ theorem Complex.uniformContinuous_ringHom_eq_id_or_conj (K : Subfield ℂ) {ψ : simp only [id, Set.image_id'] rfl ⟩ convert DenseRange.comp (Function.Surjective.denseRange _) - (DenseEmbedding.subtype denseEmbedding_id (K : Set ℂ)).dense (by continuity : Continuous j) + (IsDenseEmbedding.id.subtype (K : Set ℂ)).dense (by continuity : Continuous j) rintro ⟨y, hy⟩ use ⟨y, by diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index 1fd5e8eaa6966..22cc2d6c6f682 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -33,11 +33,14 @@ theorem uniformContinuous_coe_real : UniformContinuous ((↑) : ℚ → ℝ) := theorem uniformEmbedding_coe_real : UniformEmbedding ((↑) : ℚ → ℝ) := uniformEmbedding_comap Rat.cast_injective -theorem denseEmbedding_coe_real : DenseEmbedding ((↑) : ℚ → ℝ) := - uniformEmbedding_coe_real.denseEmbedding Rat.denseRange_cast +theorem isDenseEmbedding_coe_real : IsDenseEmbedding ((↑) : ℚ → ℝ) := + uniformEmbedding_coe_real.isDenseEmbedding Rat.denseRange_cast + +@[deprecated (since := "2024-09-30")] +alias denseEmbedding_coe_real := isDenseEmbedding_coe_real theorem embedding_coe_real : Embedding ((↑) : ℚ → ℝ) := - denseEmbedding_coe_real.to_embedding + isDenseEmbedding_coe_real.to_embedding theorem continuous_coe_real : Continuous ((↑) : ℚ → ℝ) := uniformContinuous_coe_real.continuous diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index 16c498e04a27f..69619e025121f 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -39,7 +39,7 @@ namespace Rat variable {p q : ℚ} {s t : Set ℚ} theorem interior_compact_eq_empty (hs : IsCompact s) : interior s = ∅ := - denseEmbedding_coe_real.toIsDenseInducing.interior_compact_eq_empty dense_irrational hs + isDenseEmbedding_coe_real.toIsDenseInducing.interior_compact_eq_empty dense_irrational hs theorem dense_compl_compact (hs : IsCompact s) : Dense sᶜ := interior_eq_empty_iff_dense_compl.1 (interior_compact_eq_empty hs) diff --git a/Mathlib/Topology/Instances/RealVectorSpace.lean b/Mathlib/Topology/Instances/RealVectorSpace.lean index a7eaa801c1d9f..5d4327d792021 100644 --- a/Mathlib/Topology/Instances/RealVectorSpace.lean +++ b/Mathlib/Topology/Instances/RealVectorSpace.lean @@ -23,7 +23,7 @@ theorem map_real_smul {G} [FunLike G E F] [AddMonoidHomClass G E F] (f : G) (hf (c : ℝ) (x : E) : f (c • x) = c • f x := suffices (fun c : ℝ => f (c • x)) = fun c : ℝ => c • f x from congr_fun this c - Rat.denseEmbedding_coe_real.dense.equalizer (hf.comp <| continuous_id.smul continuous_const) + Rat.isDenseEmbedding_coe_real.dense.equalizer (hf.comp <| continuous_id.smul continuous_const) (continuous_id.smul continuous_const) (funext fun r => map_ratCast_smul f ℝ ℝ r x) namespace AddMonoidHom diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index 240f69c279ec2..3beed85e1b698 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -154,10 +154,13 @@ theorem isDenseInducing_pure : @IsDenseInducing _ _ ⊥ _ (pure : α → Ultrafi -- The following refined version will never be used /-- `pure : α → Ultrafilter α` defines a dense embedding of `α` in `Ultrafilter α`. -/ -theorem denseEmbedding_pure : @DenseEmbedding _ _ ⊥ _ (pure : α → Ultrafilter α) := +theorem isDenseEmbedding_pure : @IsDenseEmbedding _ _ ⊥ _ (pure : α → Ultrafilter α) := letI : TopologicalSpace α := ⊥ { isDenseInducing_pure with inj := ultrafilter_pure_injective } +@[deprecated (since := "2024-09-30")] +alias denseEmbedding_pure := isDenseEmbedding_pure + end Embedding section Extension diff --git a/Mathlib/Topology/UniformSpace/CompareReals.lean b/Mathlib/Topology/UniformSpace/CompareReals.lean index 709fc27902447..7a90e73b4afcc 100644 --- a/Mathlib/Topology/UniformSpace/CompareReals.lean +++ b/Mathlib/Topology/UniformSpace/CompareReals.lean @@ -73,7 +73,7 @@ def rationalCauSeqPkg : @AbstractCompletion ℚ <| (@AbsoluteValue.abs ℚ _).un (uniformInducing := by rw [Rat.uniformSpace_eq] exact Rat.uniformEmbedding_coe_real.toUniformInducing) - (dense := Rat.denseEmbedding_coe_real.dense) + (dense := Rat.isDenseEmbedding_coe_real.dense) namespace CompareReals diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index b038f0744d751..0f386364453d4 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -183,12 +183,15 @@ theorem denseRange_pureCauchy : DenseRange (pureCauchy : α → CauchyFilter α) theorem isDenseInducing_pureCauchy : IsDenseInducing (pureCauchy : α → CauchyFilter α) := uniformInducing_pureCauchy.isDenseInducing denseRange_pureCauchy -theorem denseEmbedding_pureCauchy : DenseEmbedding (pureCauchy : α → CauchyFilter α) := - uniformEmbedding_pureCauchy.denseEmbedding denseRange_pureCauchy +theorem isDenseEmbedding_pureCauchy : IsDenseEmbedding (pureCauchy : α → CauchyFilter α) := + uniformEmbedding_pureCauchy.isDenseEmbedding denseRange_pureCauchy + +@[deprecated (since := "2024-09-30")] +alias denseEmbedding_pureCauchy := isDenseEmbedding_pureCauchy theorem nonempty_cauchyFilter_iff : Nonempty (CauchyFilter α) ↔ Nonempty α := by constructor <;> rintro ⟨c⟩ - · have := eq_univ_iff_forall.1 denseEmbedding_pureCauchy.toIsDenseInducing.closure_range c + · have := eq_univ_iff_forall.1 isDenseEmbedding_pureCauchy.toIsDenseInducing.closure_range c obtain ⟨_, ⟨_, a, _⟩⟩ := mem_closure_iff.1 this _ isOpen_univ trivial exact ⟨a⟩ · exact ⟨pureCauchy c⟩ @@ -377,9 +380,12 @@ open TopologicalSpace instance separableSpace_completion [SeparableSpace α] : SeparableSpace (Completion α) := Completion.isDenseInducing_coe.separableSpace -theorem denseEmbedding_coe [T0Space α] : DenseEmbedding ((↑) : α → Completion α) := +theorem isDenseEmbedding_coe [T0Space α] : IsDenseEmbedding ((↑) : α → Completion α) := { isDenseInducing_coe with inj := separated_pureCauchy_injective } +@[deprecated (since := "2024-09-30")] +alias denseEmbedding_coe := isDenseEmbedding_coe + theorem denseRange_coe₂ : DenseRange fun x : α × β => ((x.1 : Completion α), (x.2 : Completion β)) := denseRange_coe.prod_map denseRange_coe diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 9029d071993e6..048ea735a5df4 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -219,10 +219,13 @@ protected theorem UniformEmbedding.embedding {f : α → β} (h : UniformEmbeddi { toInducing := h.toUniformInducing.inducing inj := h.inj } -theorem UniformEmbedding.denseEmbedding {f : α → β} (h : UniformEmbedding f) (hd : DenseRange f) : - DenseEmbedding f := +theorem UniformEmbedding.isDenseEmbedding {f : α → β} (h : UniformEmbedding f) (hd : DenseRange f) : + IsDenseEmbedding f := { h.embedding with dense := hd } +@[deprecated (since := "2024-09-30")] +alias UniformEmbedding.denseEmbedding := UniformEmbedding.isDenseEmbedding + theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] [T0Space β] {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) (hf : Pairwise fun x y => (f x, f y) ∉ s) : ClosedEmbedding f := by @@ -246,9 +249,9 @@ theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α exact ⟨e x, hxV, mem_image_of_mem e hxU⟩ theorem uniformEmbedding_subtypeEmb (p : α → Prop) {e : α → β} (ue : UniformEmbedding e) - (de : DenseEmbedding e) : UniformEmbedding (DenseEmbedding.subtypeEmb p e) := + (de : IsDenseEmbedding e) : UniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := { comap_uniformity := by - simp [comap_comap, Function.comp_def, DenseEmbedding.subtypeEmb, uniformity_subtype, + simp [comap_comap, Function.comp_def, IsDenseEmbedding.subtypeEmb, uniformity_subtype, ue.comap_uniformity.symm] inj := (de.subtype p).inj } @@ -434,14 +437,15 @@ theorem uniform_extend_subtype [CompleteSpace γ] {p : α → Prop} {e : α → {s : Set α} (hf : UniformContinuous fun x : Subtype p => f x.val) (he : UniformEmbedding e) (hd : ∀ x : β, x ∈ closure (range e)) (hb : closure (e '' s) ∈ 𝓝 b) (hs : IsClosed s) (hp : ∀ x ∈ s, p x) : ∃ c, Tendsto f (comap e (𝓝 b)) (𝓝 c) := by - have de : DenseEmbedding e := he.denseEmbedding hd - have de' : DenseEmbedding (DenseEmbedding.subtypeEmb p e) := de.subtype p - have ue' : UniformEmbedding (DenseEmbedding.subtypeEmb p e) := uniformEmbedding_subtypeEmb _ he de + have de : IsDenseEmbedding e := he.isDenseEmbedding hd + have de' : IsDenseEmbedding (IsDenseEmbedding.subtypeEmb p e) := de.subtype p + have ue' : UniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := + uniformEmbedding_subtypeEmb _ he de have : b ∈ closure (e '' { x | p x }) := (closure_mono <| monotone_image <| hp) (mem_of_mem_nhds hb) let ⟨c, hc⟩ := uniformly_extend_exists ue'.toUniformInducing de'.dense hf ⟨b, this⟩ replace hc : Tendsto (f ∘ Subtype.val (p := p)) (((𝓝 b).comap e).comap Subtype.val) (𝓝 c) := by - simpa only [nhds_subtype_eq_comap, comap_comap, DenseEmbedding.subtypeEmb_coe] using hc + simpa only [nhds_subtype_eq_comap, comap_comap, IsDenseEmbedding.subtypeEmb_coe] using hc refine ⟨c, (tendsto_comap'_iff ?_).1 hc⟩ rw [Subtype.range_coe_subtype] exact ⟨_, hb, by rwa [← de.toInducing.closure_eq_preimage_closure_image, hs.closure_eq]⟩ diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index e8be7c1b7ded7..d66cd78fe7c09 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -998,7 +998,7 @@ decide_True' DedekindDomain.ProdAdicCompletions.algebra' DedekindDomain.ProdAdicCompletions.algebraMap_apply' DedekindDomain.ProdAdicCompletions.IsFiniteAdele.algebraMap' -DenseEmbedding.mk' +IsDenseEmbedding.mk' Dense.exists_ge' Dense.exists_le' IsDenseInducing.extend_eq_at' From cb03ac23046975efec06f8795593e438d39a373b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 14:37:27 +0000 Subject: [PATCH 064/174] feat: Left/right commutativity of `min`/`max` on `Nat` (#17287) These will be used to prove `List` lemmas that can't afford to import the more general statements From LeanCamCombi --- Mathlib/Data/Nat/Defs.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index b5e07ed1bde05..81287ae2f16f5 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -141,6 +141,18 @@ lemma le_one_iff_eq_zero_or_eq_one : ∀ {n : ℕ}, n ≤ 1 ↔ n = 0 ∨ n = 1 lemma one_le_of_lt (h : a < b) : 1 ≤ b := Nat.lt_of_le_of_lt (Nat.zero_le _) h +protected lemma min_left_comm (a b c : ℕ) : min a (min b c) = min b (min a c) := by + rw [← Nat.min_assoc, ← Nat.min_assoc, b.min_comm] + +protected lemma max_left_comm (a b c : ℕ) : max a (max b c) = max b (max a c) := by + rw [← Nat.max_assoc, ← Nat.max_assoc, b.max_comm] + +protected lemma min_right_comm (a b c : ℕ) : min (min a b) c = min (min a c) b := by + rw [Nat.min_assoc, Nat.min_assoc, b.min_comm] + +protected lemma max_right_comm (a b c : ℕ) : max (max a b) c = max (max a c) b := by + rw [Nat.max_assoc, Nat.max_assoc, b.max_comm] + @[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := by omega @[simp] lemma max_eq_zero_iff : max m n = 0 ↔ m = 0 ∧ n = 0 := by omega From a6e5517e7b3a74624e836bb636733bb46d970f0f Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Mon, 30 Sep 2024 15:34:53 +0000 Subject: [PATCH 065/174] chore(AlgebraicGeometry/EllipticCurve/*): split `VariableChange` into a new file (#17199) --- Mathlib.lean | 1 + .../EllipticCurve/Affine.lean | 1 + .../EllipticCurve/VariableChange.lean | 325 ++++++++++++++++++ .../EllipticCurve/Weierstrass.lean | 269 --------------- 4 files changed, 327 insertions(+), 269 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean diff --git a/Mathlib.lean b/Mathlib.lean index 66017fa31543d..71ccbdfe6ec3c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -833,6 +833,7 @@ import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Degree import Mathlib.AlgebraicGeometry.EllipticCurve.Group import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian import Mathlib.AlgebraicGeometry.EllipticCurve.Projective +import Mathlib.AlgebraicGeometry.EllipticCurve.VariableChange import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass import Mathlib.AlgebraicGeometry.FunctionField import Mathlib.AlgebraicGeometry.GammaSpecAdjunction diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 8e25ace403b09..d3cc6dc2264fe 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -5,6 +5,7 @@ Authors: David Kurniadi Angdinata -/ import Mathlib.Algebra.Polynomial.Bivariate import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass +import Mathlib.AlgebraicGeometry.EllipticCurve.VariableChange /-! # Affine coordinates for Weierstrass curves diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean new file mode 100644 index 0000000000000..8f77cea60aaad --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean @@ -0,0 +1,325 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, David Kurniadi Angdinata, Jz Pan +-/ +import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass + +/-! +# Change of variables of Weierstrass curves + +This file defines admissible linear change of variables of Weierstrass curves. + +## Main definitions + + * `WeierstrassCurve.VariableChange`: a change of variables of Weierstrass curves. + * `WeierstrassCurve.variableChange`: the Weierstrass curve induced by a change of variables. + * `WeierstrassCurve.instMulActionVariableChange`: change of variables act on Weierstrass curves. + * `EllipticCurve.variableChange`: the elliptic curve induced by a change of variables. + * `EllipticCurve.instMulActionVariableChange`: change of variables act on elliptic curves. + +## Main statements + + * `EllipticCurve.variableChange_j`: the j-invariant of an elliptic curve is invariant under an + admissible linear change of variables. + +## References + + * [J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, weierstrass equation, change of variables +-/ + +local macro "map_simp" : tactic => + `(tactic| simp only [map_ofNat, map_neg, map_add, map_sub, map_mul, map_pow]) + +universe s u v w + +namespace WeierstrassCurve + +variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) + +section VariableChange + +/-! ### Variable changes -/ + +/-- An admissible linear change of variables of Weierstrass curves defined over a ring `R` given by +a tuple $(u, r, s, t)$ for some $u \in R^\times$ and some $r, s, t \in R$. As a matrix, it is +$\begin{pmatrix} u^2 & 0 & r \cr u^2s & u^3 & t \cr 0 & 0 & 1 \end{pmatrix}$. -/ +@[ext] +structure VariableChange (R : Type u) [CommRing R] where + /-- The `u` coefficient of an admissible linear change of variables, which must be a unit. -/ + u : Rˣ + /-- The `r` coefficient of an admissible linear change of variables. -/ + r : R + /-- The `s` coefficient of an admissible linear change of variables. -/ + s : R + /-- The `t` coefficient of an admissible linear change of variables. -/ + t : R + +namespace VariableChange + +variable (C C' C'' : VariableChange R) + +/-- The identity linear change of variables given by the identity matrix. -/ +def id : VariableChange R := + ⟨1, 0, 0, 0⟩ + +/-- The composition of two linear changes of variables given by matrix multiplication. -/ +def comp : VariableChange R where + u := C.u * C'.u + r := C.r * C'.u ^ 2 + C'.r + s := C'.u * C.s + C'.s + t := C.t * C'.u ^ 3 + C.r * C'.s * C'.u ^ 2 + C'.t + +/-- The inverse of a linear change of variables given by matrix inversion. -/ +def inv : VariableChange R where + u := C.u⁻¹ + r := -C.r * C.u⁻¹ ^ 2 + s := -C.s * C.u⁻¹ + t := (C.r * C.s - C.t) * C.u⁻¹ ^ 3 + +lemma id_comp (C : VariableChange R) : comp id C = C := by + simp only [comp, id, zero_add, zero_mul, mul_zero, one_mul] + +lemma comp_id (C : VariableChange R) : comp C id = C := by + simp only [comp, id, add_zero, mul_zero, one_mul, mul_one, one_pow, Units.val_one] + +lemma comp_left_inv (C : VariableChange R) : comp (inv C) C = id := by + rw [comp, id, inv] + ext <;> dsimp only + · exact C.u.inv_mul + · linear_combination (norm := ring1) -C.r * pow_mul_pow_eq_one 2 C.u.inv_mul + · linear_combination (norm := ring1) -C.s * C.u.inv_mul + · linear_combination (norm := ring1) (C.r * C.s - C.t) * pow_mul_pow_eq_one 3 C.u.inv_mul + + -C.r * C.s * pow_mul_pow_eq_one 2 C.u.inv_mul + +lemma comp_assoc (C C' C'' : VariableChange R) : comp (comp C C') C'' = comp C (comp C' C'') := by + ext <;> simp only [comp, Units.val_mul] <;> ring1 + +instance instGroup : Group (VariableChange R) where + one := id + inv := inv + mul := comp + one_mul := id_comp + mul_one := comp_id + inv_mul_cancel := comp_left_inv + mul_assoc := comp_assoc + +end VariableChange + +variable (C : VariableChange R) + +/-- The Weierstrass curve over `R` induced by an admissible linear change of variables +$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. -/ +@[simps] +def variableChange : WeierstrassCurve R where + a₁ := C.u⁻¹ * (W.a₁ + 2 * C.s) + a₂ := C.u⁻¹ ^ 2 * (W.a₂ - C.s * W.a₁ + 3 * C.r - C.s ^ 2) + a₃ := C.u⁻¹ ^ 3 * (W.a₃ + C.r * W.a₁ + 2 * C.t) + a₄ := C.u⁻¹ ^ 4 * (W.a₄ - C.s * W.a₃ + 2 * C.r * W.a₂ - (C.t + C.r * C.s) * W.a₁ + 3 * C.r ^ 2 + - 2 * C.s * C.t) + a₆ := C.u⁻¹ ^ 6 * (W.a₆ + C.r * W.a₄ + C.r ^ 2 * W.a₂ + C.r ^ 3 - C.t * W.a₃ - C.t ^ 2 + - C.r * C.t * W.a₁) + +lemma variableChange_id : W.variableChange VariableChange.id = W := by + rw [VariableChange.id, variableChange, inv_one, Units.val_one] + ext <;> (dsimp only; ring1) + +lemma variableChange_comp (C C' : VariableChange R) (W : WeierstrassCurve R) : + W.variableChange (C.comp C') = (W.variableChange C').variableChange C := by + simp only [VariableChange.comp, variableChange] + ext <;> simp only [mul_inv, Units.val_mul] + · linear_combination (norm := ring1) ↑C.u⁻¹ * C.s * 2 * C'.u.inv_mul + · linear_combination (norm := ring1) + C.s * (-C'.s * 2 - W.a₁) * C.u⁻¹ ^ 2 * ↑C'.u⁻¹ * C'.u.inv_mul + + (C.r * 3 - C.s ^ 2) * C.u⁻¹ ^ 2 * pow_mul_pow_eq_one 2 C'.u.inv_mul + · linear_combination (norm := ring1) + C.r * (C'.s * 2 + W.a₁) * C.u⁻¹ ^ 3 * ↑C'.u⁻¹ * pow_mul_pow_eq_one 2 C'.u.inv_mul + + C.t * 2 * C.u⁻¹ ^ 3 * pow_mul_pow_eq_one 3 C'.u.inv_mul + · linear_combination (norm := ring1) + C.s * (-W.a₃ - C'.r * W.a₁ - C'.t * 2) * C.u⁻¹ ^ 4 * C'.u⁻¹ ^ 3 * C'.u.inv_mul + + C.u⁻¹ ^ 4 * C'.u⁻¹ ^ 2 * (C.r * C'.r * 6 + C.r * W.a₂ * 2 - C'.s * C.r * W.a₁ * 2 + - C'.s ^ 2 * C.r * 2) * pow_mul_pow_eq_one 2 C'.u.inv_mul + - C.u⁻¹ ^ 4 * ↑C'.u⁻¹ * (C.s * C'.s * C.r * 2 + C.s * C.r * W.a₁ + C'.s * C.t * 2 + + C.t * W.a₁) * pow_mul_pow_eq_one 3 C'.u.inv_mul + + C.u⁻¹ ^ 4 * (C.r ^ 2 * 3 - C.s * C.t * 2) * pow_mul_pow_eq_one 4 C'.u.inv_mul + · linear_combination (norm := ring1) + C.r * C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 4 * (C'.r * W.a₂ * 2 - C'.r * C'.s * W.a₁ + C'.r ^ 2 * 3 + W.a₄ + - C'.s * C'.t * 2 - C'.s * W.a₃ - C'.t * W.a₁) * pow_mul_pow_eq_one 2 C'.u.inv_mul + - C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 3 * C.t * (C'.r * W.a₁ + C'.t * 2 + W.a₃) + * pow_mul_pow_eq_one 3 C'.u.inv_mul + + C.r ^ 2 * C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 2 * (C'.r * 3 + W.a₂ - C'.s * W.a₁ - C'.s ^ 2) + * pow_mul_pow_eq_one 4 C'.u.inv_mul + - C.r * C.t * C.u⁻¹ ^ 6 * ↑C'.u⁻¹ * (C'.s * 2 + W.a₁) * pow_mul_pow_eq_one 5 C'.u.inv_mul + + C.u⁻¹ ^ 6 * (C.r ^ 3 - C.t ^ 2) * pow_mul_pow_eq_one 6 C'.u.inv_mul + +instance instMulActionVariableChange : MulAction (VariableChange R) (WeierstrassCurve R) where + smul := fun C W => W.variableChange C + one_smul := variableChange_id + mul_smul := variableChange_comp + +@[simp] +lemma variableChange_b₂ : (W.variableChange C).b₂ = C.u⁻¹ ^ 2 * (W.b₂ + 12 * C.r) := by + simp only [b₂, variableChange_a₁, variableChange_a₂] + ring1 + +@[simp] +lemma variableChange_b₄ : + (W.variableChange C).b₄ = C.u⁻¹ ^ 4 * (W.b₄ + C.r * W.b₂ + 6 * C.r ^ 2) := by + simp only [b₂, b₄, variableChange_a₁, variableChange_a₃, variableChange_a₄] + ring1 + +@[simp] +lemma variableChange_b₆ : (W.variableChange C).b₆ = + C.u⁻¹ ^ 6 * (W.b₆ + 2 * C.r * W.b₄ + C.r ^ 2 * W.b₂ + 4 * C.r ^ 3) := by + simp only [b₂, b₄, b₆, variableChange_a₃, variableChange_a₆] + ring1 + +@[simp] +lemma variableChange_b₈ : (W.variableChange C).b₈ = C.u⁻¹ ^ 8 * + (W.b₈ + 3 * C.r * W.b₆ + 3 * C.r ^ 2 * W.b₄ + C.r ^ 3 * W.b₂ + 3 * C.r ^ 4) := by + simp only [b₂, b₄, b₆, b₈, variableChange_a₁, variableChange_a₂, variableChange_a₃, + variableChange_a₄, variableChange_a₆] + ring1 + +@[simp] +lemma variableChange_c₄ : (W.variableChange C).c₄ = C.u⁻¹ ^ 4 * W.c₄ := by + simp only [c₄, variableChange_b₂, variableChange_b₄] + ring1 + +@[simp] +lemma variableChange_c₆ : (W.variableChange C).c₆ = C.u⁻¹ ^ 6 * W.c₆ := by + simp only [c₆, variableChange_b₂, variableChange_b₄, variableChange_b₆] + ring1 + +@[simp] +lemma variableChange_Δ : (W.variableChange C).Δ = C.u⁻¹ ^ 12 * W.Δ := by + simp only [b₂, b₄, b₆, b₈, Δ, variableChange_a₁, variableChange_a₂, variableChange_a₃, + variableChange_a₄, variableChange_a₆] + ring1 + +end VariableChange + +section BaseChange + +/-! ### Maps and base changes of variable changes -/ + +variable {A : Type v} [CommRing A] (φ : R →+* A) + +namespace VariableChange + +variable (C : VariableChange R) + +/-- The change of variables mapped over a ring homomorphism `φ : R →+* A`. -/ +@[simps] +def map : VariableChange A := + ⟨Units.map φ C.u, φ C.r, φ C.s, φ C.t⟩ + +variable (A) + +/-- The change of variables base changed to an algebra `A` over `R`. -/ +abbrev baseChange [Algebra R A] : VariableChange A := + C.map <| algebraMap R A + +variable {A} + +@[simp] +lemma map_id : C.map (RingHom.id R) = C := + rfl + +lemma map_map {A : Type v} [CommRing A] (φ : R →+* A) {B : Type w} [CommRing B] (ψ : A →+* B) : + (C.map φ).map ψ = C.map (ψ.comp φ) := + rfl + +@[simp] +lemma map_baseChange {S : Type s} [CommRing S] [Algebra R S] {A : Type v} [CommRing A] [Algebra R A] + [Algebra S A] [IsScalarTower R S A] {B : Type w} [CommRing B] [Algebra R B] [Algebra S B] + [IsScalarTower R S B] (ψ : A →ₐ[S] B) : (C.baseChange A).map ψ = C.baseChange B := + congr_arg C.map <| ψ.comp_algebraMap_of_tower R + +lemma map_injective {φ : R →+* A} (hφ : Function.Injective φ) : + Function.Injective <| map (φ := φ) := fun _ _ h => by + rcases mk.inj h with ⟨h, _, _, _⟩ + replace h := (Units.mk.inj h).left + ext <;> apply_fun _ using hφ <;> assumption + +private lemma id_map : (id : VariableChange R).map φ = id := by + simp only [id, map] + ext <;> simp only [map_one, Units.val_one, map_zero] + +private lemma comp_map (C' : VariableChange R) : (C.comp C').map φ = (C.map φ).comp (C'.map φ) := by + simp only [comp, map] + ext <;> map_simp <;> simp only [Units.coe_map, Units.coe_map_inv, MonoidHom.coe_coe] + +/-- The map over a ring homomorphism of a change of variables is a group homomorphism. -/ +def mapHom : VariableChange R →* VariableChange A where + toFun := map φ + map_one' := id_map φ + map_mul' := comp_map φ + +end VariableChange + +lemma map_variableChange (C : VariableChange R) : + (W.map φ).variableChange (C.map φ) = (W.variableChange C).map φ := by + simp only [map, variableChange, VariableChange.map] + ext <;> map_simp <;> simp only [Units.coe_map, Units.coe_map_inv, MonoidHom.coe_coe] + +end BaseChange + +end WeierstrassCurve + +/-! ## Variable changes of elliptic curves -/ + +namespace EllipticCurve + +variable {R : Type u} [CommRing R] + +variable (E : EllipticCurve R) + +section VariableChange + +variable (C : WeierstrassCurve.VariableChange R) + +-- Porting note: was just `@[simps]` +/-- The elliptic curve over `R` induced by an admissible linear change of variables +$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. +When `R` is a field, any two Weierstrass equations isomorphic to `E` are related by this. -/ +@[simps (config := { rhsMd := .default }) a₁ a₂ a₃ a₄ a₆ Δ' toWeierstrassCurve] +def variableChange : EllipticCurve R := + ⟨E.toWeierstrassCurve.variableChange C, C.u⁻¹ ^ 12 * E.Δ', by + rw [Units.val_mul, Units.val_pow_eq_pow_val, coe_Δ', E.variableChange_Δ]⟩ + +lemma variableChange_id : E.variableChange WeierstrassCurve.VariableChange.id = E := by + simp only [variableChange, WeierstrassCurve.variableChange_id] + simp only [WeierstrassCurve.VariableChange.id, inv_one, one_pow, one_mul] + +lemma variableChange_comp (C C' : WeierstrassCurve.VariableChange R) (E : EllipticCurve R) : + E.variableChange (C.comp C') = (E.variableChange C').variableChange C := by + simp only [variableChange, WeierstrassCurve.variableChange_comp] + simp only [WeierstrassCurve.VariableChange.comp, mul_inv, mul_pow, ← mul_assoc] + +instance instMulActionVariableChange : + MulAction (WeierstrassCurve.VariableChange R) (EllipticCurve R) where + smul := fun C E => E.variableChange C + one_smul := variableChange_id + mul_smul := variableChange_comp + +lemma coe_variableChange_Δ' : (E.variableChange C).Δ' = C.u⁻¹ ^ 12 * E.Δ' := + rfl + +lemma coe_inv_variableChange_Δ' : (E.variableChange C).Δ'⁻¹ = C.u ^ 12 * E.Δ'⁻¹ := by + rw [variableChange_Δ', mul_inv, inv_pow, inv_inv] + +@[simp] +lemma variableChange_j : (E.variableChange C).j = E.j := by + rw [j, coe_inv_variableChange_Δ', Units.val_mul, Units.val_pow_eq_pow_val, + variableChange_toWeierstrassCurve, WeierstrassCurve.variableChange_c₄] + have hu : (C.u * C.u⁻¹ : R) ^ 12 = 1 := by rw [C.u.mul_inv, one_pow] + linear_combination (norm := (rw [j]; ring1)) E.j * hu + +end VariableChange + +end EllipticCurve diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index a1e64d399e356..72b395c19ab00 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -33,8 +33,6 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi * `WeierstrassCurve.ofJ0`: a Weierstrass curve whose j-invariant is 0. * `WeierstrassCurve.ofJ1728`: a Weierstrass curve whose j-invariant is 1728. * `WeierstrassCurve.ofJ`: a Weierstrass curve whose j-invariant is neither 0 nor 1728. - * `WeierstrassCurve.VariableChange`: a change of variables of Weierstrass curves. - * `WeierstrassCurve.variableChange`: the Weierstrass curve induced by a change of variables. * `WeierstrassCurve.map`: the Weierstrass curve mapped over a ring homomorphism. * `WeierstrassCurve.twoTorsionPolynomial`: the 2-torsion polynomial of a Weierstrass curve. * `EllipticCurve`: an elliptic curve over a commutative ring. @@ -48,8 +46,6 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi * `WeierstrassCurve.twoTorsionPolynomial_disc`: the discriminant of a Weierstrass curve is a constant factor of the cubic discriminant of its 2-torsion polynomial. - * `EllipticCurve.variableChange_j`: the j-invariant of an elliptic curve is invariant under an - admissible linear change of variables. * `EllipticCurve.ofJ_j`: the j-invariant of `EllipticCurve.ofJ` is equal to j. ## Implementation notes @@ -154,168 +150,6 @@ lemma c_relation : 1728 * W.Δ = W.c₄ ^ 3 - W.c₆ ^ 2 := by end Quantity -section VariableChange - -/-! ### Variable changes -/ - -/-- An admissible linear change of variables of Weierstrass curves defined over a ring `R` given by -a tuple $(u, r, s, t)$ for some $u \in R^\times$ and some $r, s, t \in R$. As a matrix, it is -$\begin{pmatrix} u^2 & 0 & r \cr u^2s & u^3 & t \cr 0 & 0 & 1 \end{pmatrix}$. -/ -@[ext] -structure VariableChange (R : Type u) [CommRing R] where - /-- The `u` coefficient of an admissible linear change of variables, which must be a unit. -/ - u : Rˣ - /-- The `r` coefficient of an admissible linear change of variables. -/ - r : R - /-- The `s` coefficient of an admissible linear change of variables. -/ - s : R - /-- The `t` coefficient of an admissible linear change of variables. -/ - t : R - -namespace VariableChange - -variable (C C' C'' : VariableChange R) - -/-- The identity linear change of variables given by the identity matrix. -/ -def id : VariableChange R := - ⟨1, 0, 0, 0⟩ - -/-- The composition of two linear changes of variables given by matrix multiplication. -/ -def comp : VariableChange R where - u := C.u * C'.u - r := C.r * C'.u ^ 2 + C'.r - s := C'.u * C.s + C'.s - t := C.t * C'.u ^ 3 + C.r * C'.s * C'.u ^ 2 + C'.t - -/-- The inverse of a linear change of variables given by matrix inversion. -/ -def inv : VariableChange R where - u := C.u⁻¹ - r := -C.r * C.u⁻¹ ^ 2 - s := -C.s * C.u⁻¹ - t := (C.r * C.s - C.t) * C.u⁻¹ ^ 3 - -lemma id_comp (C : VariableChange R) : comp id C = C := by - simp only [comp, id, zero_add, zero_mul, mul_zero, one_mul] - -lemma comp_id (C : VariableChange R) : comp C id = C := by - simp only [comp, id, add_zero, mul_zero, one_mul, mul_one, one_pow, Units.val_one] - -lemma comp_left_inv (C : VariableChange R) : comp (inv C) C = id := by - rw [comp, id, inv] - ext <;> dsimp only - · exact C.u.inv_mul - · linear_combination (norm := ring1) -C.r * pow_mul_pow_eq_one 2 C.u.inv_mul - · linear_combination (norm := ring1) -C.s * C.u.inv_mul - · linear_combination (norm := ring1) (C.r * C.s - C.t) * pow_mul_pow_eq_one 3 C.u.inv_mul - + -C.r * C.s * pow_mul_pow_eq_one 2 C.u.inv_mul - -lemma comp_assoc (C C' C'' : VariableChange R) : comp (comp C C') C'' = comp C (comp C' C'') := by - ext <;> simp only [comp, Units.val_mul] <;> ring1 - -instance instGroup : Group (VariableChange R) where - one := id - inv := inv - mul := comp - one_mul := id_comp - mul_one := comp_id - inv_mul_cancel := comp_left_inv - mul_assoc := comp_assoc - -end VariableChange - -variable (C : VariableChange R) - -/-- The Weierstrass curve over `R` induced by an admissible linear change of variables -$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. -/ -@[simps] -def variableChange : WeierstrassCurve R where - a₁ := C.u⁻¹ * (W.a₁ + 2 * C.s) - a₂ := C.u⁻¹ ^ 2 * (W.a₂ - C.s * W.a₁ + 3 * C.r - C.s ^ 2) - a₃ := C.u⁻¹ ^ 3 * (W.a₃ + C.r * W.a₁ + 2 * C.t) - a₄ := C.u⁻¹ ^ 4 * (W.a₄ - C.s * W.a₃ + 2 * C.r * W.a₂ - (C.t + C.r * C.s) * W.a₁ + 3 * C.r ^ 2 - - 2 * C.s * C.t) - a₆ := C.u⁻¹ ^ 6 * (W.a₆ + C.r * W.a₄ + C.r ^ 2 * W.a₂ + C.r ^ 3 - C.t * W.a₃ - C.t ^ 2 - - C.r * C.t * W.a₁) - -lemma variableChange_id : W.variableChange VariableChange.id = W := by - rw [VariableChange.id, variableChange, inv_one, Units.val_one] - ext <;> (dsimp only; ring1) - -lemma variableChange_comp (C C' : VariableChange R) (W : WeierstrassCurve R) : - W.variableChange (C.comp C') = (W.variableChange C').variableChange C := by - simp only [VariableChange.comp, variableChange] - ext <;> simp only [mul_inv, Units.val_mul] - · linear_combination (norm := ring1) ↑C.u⁻¹ * C.s * 2 * C'.u.inv_mul - · linear_combination (norm := ring1) - C.s * (-C'.s * 2 - W.a₁) * C.u⁻¹ ^ 2 * ↑C'.u⁻¹ * C'.u.inv_mul - + (C.r * 3 - C.s ^ 2) * C.u⁻¹ ^ 2 * pow_mul_pow_eq_one 2 C'.u.inv_mul - · linear_combination (norm := ring1) - C.r * (C'.s * 2 + W.a₁) * C.u⁻¹ ^ 3 * ↑C'.u⁻¹ * pow_mul_pow_eq_one 2 C'.u.inv_mul - + C.t * 2 * C.u⁻¹ ^ 3 * pow_mul_pow_eq_one 3 C'.u.inv_mul - · linear_combination (norm := ring1) - C.s * (-W.a₃ - C'.r * W.a₁ - C'.t * 2) * C.u⁻¹ ^ 4 * C'.u⁻¹ ^ 3 * C'.u.inv_mul - + C.u⁻¹ ^ 4 * C'.u⁻¹ ^ 2 * (C.r * C'.r * 6 + C.r * W.a₂ * 2 - C'.s * C.r * W.a₁ * 2 - - C'.s ^ 2 * C.r * 2) * pow_mul_pow_eq_one 2 C'.u.inv_mul - - C.u⁻¹ ^ 4 * ↑C'.u⁻¹ * (C.s * C'.s * C.r * 2 + C.s * C.r * W.a₁ + C'.s * C.t * 2 - + C.t * W.a₁) * pow_mul_pow_eq_one 3 C'.u.inv_mul - + C.u⁻¹ ^ 4 * (C.r ^ 2 * 3 - C.s * C.t * 2) * pow_mul_pow_eq_one 4 C'.u.inv_mul - · linear_combination (norm := ring1) - C.r * C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 4 * (C'.r * W.a₂ * 2 - C'.r * C'.s * W.a₁ + C'.r ^ 2 * 3 + W.a₄ - - C'.s * C'.t * 2 - C'.s * W.a₃ - C'.t * W.a₁) * pow_mul_pow_eq_one 2 C'.u.inv_mul - - C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 3 * C.t * (C'.r * W.a₁ + C'.t * 2 + W.a₃) - * pow_mul_pow_eq_one 3 C'.u.inv_mul - + C.r ^ 2 * C.u⁻¹ ^ 6 * C'.u⁻¹ ^ 2 * (C'.r * 3 + W.a₂ - C'.s * W.a₁ - C'.s ^ 2) - * pow_mul_pow_eq_one 4 C'.u.inv_mul - - C.r * C.t * C.u⁻¹ ^ 6 * ↑C'.u⁻¹ * (C'.s * 2 + W.a₁) * pow_mul_pow_eq_one 5 C'.u.inv_mul - + C.u⁻¹ ^ 6 * (C.r ^ 3 - C.t ^ 2) * pow_mul_pow_eq_one 6 C'.u.inv_mul - -instance instMulActionVariableChange : MulAction (VariableChange R) (WeierstrassCurve R) where - smul := fun C W => W.variableChange C - one_smul := variableChange_id - mul_smul := variableChange_comp - -@[simp] -lemma variableChange_b₂ : (W.variableChange C).b₂ = C.u⁻¹ ^ 2 * (W.b₂ + 12 * C.r) := by - simp only [b₂, variableChange_a₁, variableChange_a₂] - ring1 - -@[simp] -lemma variableChange_b₄ : - (W.variableChange C).b₄ = C.u⁻¹ ^ 4 * (W.b₄ + C.r * W.b₂ + 6 * C.r ^ 2) := by - simp only [b₂, b₄, variableChange_a₁, variableChange_a₃, variableChange_a₄] - ring1 - -@[simp] -lemma variableChange_b₆ : (W.variableChange C).b₆ = - C.u⁻¹ ^ 6 * (W.b₆ + 2 * C.r * W.b₄ + C.r ^ 2 * W.b₂ + 4 * C.r ^ 3) := by - simp only [b₂, b₄, b₆, variableChange_a₃, variableChange_a₆] - ring1 - -@[simp] -lemma variableChange_b₈ : (W.variableChange C).b₈ = C.u⁻¹ ^ 8 * - (W.b₈ + 3 * C.r * W.b₆ + 3 * C.r ^ 2 * W.b₄ + C.r ^ 3 * W.b₂ + 3 * C.r ^ 4) := by - simp only [b₂, b₄, b₆, b₈, variableChange_a₁, variableChange_a₂, variableChange_a₃, - variableChange_a₄, variableChange_a₆] - ring1 - -@[simp] -lemma variableChange_c₄ : (W.variableChange C).c₄ = C.u⁻¹ ^ 4 * W.c₄ := by - simp only [c₄, variableChange_b₂, variableChange_b₄] - ring1 - -@[simp] -lemma variableChange_c₆ : (W.variableChange C).c₆ = C.u⁻¹ ^ 6 * W.c₆ := by - simp only [c₆, variableChange_b₂, variableChange_b₄, variableChange_b₆] - ring1 - -@[simp] -lemma variableChange_Δ : (W.variableChange C).Δ = C.u⁻¹ ^ 12 * W.Δ := by - simp only [b₂, b₄, b₆, b₈, Δ, variableChange_a₁, variableChange_a₂, variableChange_a₃, - variableChange_a₄, variableChange_a₆] - ring1 - -end VariableChange - section BaseChange /-! ### Maps and base changes -/ @@ -388,64 +222,6 @@ lemma map_injective {φ : R →+* A} (hφ : Function.Injective φ) : rcases mk.inj h with ⟨_, _, _, _, _⟩ ext <;> apply_fun _ using hφ <;> assumption -namespace VariableChange - -variable (C : VariableChange R) - -/-- The change of variables mapped over a ring homomorphism `φ : R →+* A`. -/ -@[simps] -def map : VariableChange A := - ⟨Units.map φ C.u, φ C.r, φ C.s, φ C.t⟩ - -variable (A) - -/-- The change of variables base changed to an algebra `A` over `R`. -/ -abbrev baseChange [Algebra R A] : VariableChange A := - C.map <| algebraMap R A - -variable {A} - -@[simp] -lemma map_id : C.map (RingHom.id R) = C := - rfl - -lemma map_map {A : Type v} [CommRing A] (φ : R →+* A) {B : Type w} [CommRing B] (ψ : A →+* B) : - (C.map φ).map ψ = C.map (ψ.comp φ) := - rfl - -@[simp] -lemma map_baseChange {S : Type s} [CommRing S] [Algebra R S] {A : Type v} [CommRing A] [Algebra R A] - [Algebra S A] [IsScalarTower R S A] {B : Type w} [CommRing B] [Algebra R B] [Algebra S B] - [IsScalarTower R S B] (ψ : A →ₐ[S] B) : (C.baseChange A).map ψ = C.baseChange B := - congr_arg C.map <| ψ.comp_algebraMap_of_tower R - -lemma map_injective {φ : R →+* A} (hφ : Function.Injective φ) : - Function.Injective <| map (φ := φ) := fun _ _ h => by - rcases mk.inj h with ⟨h, _, _, _⟩ - replace h := (Units.mk.inj h).left - ext <;> apply_fun _ using hφ <;> assumption - -private lemma id_map : (id : VariableChange R).map φ = id := by - simp only [id, map] - ext <;> simp only [map_one, Units.val_one, map_zero] - -private lemma comp_map (C' : VariableChange R) : (C.comp C').map φ = (C.map φ).comp (C'.map φ) := by - simp only [comp, map] - ext <;> map_simp <;> simp only [Units.coe_map, Units.coe_map_inv, MonoidHom.coe_coe] - -/-- The map over a ring homomorphism of a change of variables is a group homomorphism. -/ -def mapHom : VariableChange R →* VariableChange A where - toFun := map φ - map_one' := id_map φ - map_mul' := comp_map φ - -end VariableChange - -lemma map_variableChange (C : VariableChange R) : - (W.map φ).variableChange (C.map φ) = (W.variableChange C).map φ := by - simp only [map, variableChange, VariableChange.map] - ext <;> map_simp <;> simp only [Units.coe_map, Units.coe_map_inv, MonoidHom.coe_coe] - end BaseChange section TorsionPolynomial @@ -568,51 +344,6 @@ lemma twoTorsionPolynomial_disc_ne_zero [Nontrivial R] [Invertible (2 : R)] : E.twoTorsionPolynomial.disc ≠ 0 := E.toWeierstrassCurve.twoTorsionPolynomial_disc_ne_zero <| E.coe_Δ' ▸ E.Δ'.isUnit -section VariableChange - -/-! ### Variable changes -/ - -variable (C : WeierstrassCurve.VariableChange R) - --- Porting note: was just `@[simps]` -/-- The elliptic curve over `R` induced by an admissible linear change of variables -$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. -When `R` is a field, any two Weierstrass equations isomorphic to `E` are related by this. -/ -@[simps (config := { rhsMd := .default }) a₁ a₂ a₃ a₄ a₆ Δ' toWeierstrassCurve] -def variableChange : EllipticCurve R := - ⟨E.toWeierstrassCurve.variableChange C, C.u⁻¹ ^ 12 * E.Δ', by - rw [Units.val_mul, Units.val_pow_eq_pow_val, coe_Δ', E.variableChange_Δ]⟩ - -lemma variableChange_id : E.variableChange WeierstrassCurve.VariableChange.id = E := by - simp only [variableChange, WeierstrassCurve.variableChange_id] - simp only [WeierstrassCurve.VariableChange.id, inv_one, one_pow, one_mul] - -lemma variableChange_comp (C C' : WeierstrassCurve.VariableChange R) (E : EllipticCurve R) : - E.variableChange (C.comp C') = (E.variableChange C').variableChange C := by - simp only [variableChange, WeierstrassCurve.variableChange_comp] - simp only [WeierstrassCurve.VariableChange.comp, mul_inv, mul_pow, ← mul_assoc] - -instance instMulActionVariableChange : - MulAction (WeierstrassCurve.VariableChange R) (EllipticCurve R) where - smul := fun C E => E.variableChange C - one_smul := variableChange_id - mul_smul := variableChange_comp - -lemma coe_variableChange_Δ' : (E.variableChange C).Δ' = C.u⁻¹ ^ 12 * E.Δ' := - rfl - -lemma coe_inv_variableChange_Δ' : (E.variableChange C).Δ'⁻¹ = C.u ^ 12 * E.Δ'⁻¹ := by - rw [variableChange_Δ', mul_inv, inv_pow, inv_inv] - -@[simp] -lemma variableChange_j : (E.variableChange C).j = E.j := by - rw [j, coe_inv_variableChange_Δ', Units.val_mul, Units.val_pow_eq_pow_val, - variableChange_toWeierstrassCurve, WeierstrassCurve.variableChange_c₄] - have hu : (C.u * C.u⁻¹ : R) ^ 12 = 1 := by rw [C.u.mul_inv, one_pow] - linear_combination (norm := (rw [j]; ring1)) E.j * hu - -end VariableChange - section BaseChange /-! ### Maps and base changes -/ From 656765e971aeaab64e350a31d01d2e774c2bb0c8 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Mon, 30 Sep 2024 15:50:42 +0000 Subject: [PATCH 066/174] feat(MetricSpace/Ultra): `TotallyDisconnectedSpace X` when X is ultrametric (#17297) --- Mathlib.lean | 1 + .../Ultra/TotallyDisconnected.lean | 25 +++++++++++++++++++ 2 files changed, 26 insertions(+) create mode 100644 Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean diff --git a/Mathlib.lean b/Mathlib.lean index 71ccbdfe6ec3c..8211ab834a536 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4711,6 +4711,7 @@ import Mathlib.Topology.MetricSpace.ShrinkingLemma import Mathlib.Topology.MetricSpace.ThickenedIndicator import Mathlib.Topology.MetricSpace.Thickening import Mathlib.Topology.MetricSpace.Ultra.Basic +import Mathlib.Topology.MetricSpace.Ultra.TotallyDisconnected import Mathlib.Topology.Metrizable.Basic import Mathlib.Topology.Metrizable.ContinuousMap import Mathlib.Topology.Metrizable.Uniformity diff --git a/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean b/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean new file mode 100644 index 0000000000000..161f861120e27 --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Ultra/TotallyDisconnected.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2024 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky, David Loeffler +-/ +import Mathlib.Topology.MetricSpace.Defs +import Mathlib.Topology.MetricSpace.Ultra.Basic + +/-! +# Ultrametric spaces are totally disconnected + +In a metric space with an ultrametric, the space is totally disconnected. + +## Tags + +ultrametric, nonarchimedean, totally disconnected +-/ +open Metric IsUltrametricDist + +instance {X : Type*} [MetricSpace X] [IsUltrametricDist X] : TotallyDisconnectedSpace X := by + refine (totallyDisconnectedSpace_iff X).mpr (isTotallyDisconnected_of_isClopen_set fun x y h ↦ ?_) + obtain ⟨r, hr, hr'⟩ := exists_between (dist_pos.mpr h) + refine ⟨_, IsUltrametricDist.isClopen_ball x r, ?_, ?_⟩ + · simp only [mem_ball, dist_self, hr] + · simp only [mem_ball, dist_comm, not_lt, hr'.le] From 0b5e485b17d5b7b9666011202798fb5f62b5c8aa Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 30 Sep 2024 16:44:19 +0000 Subject: [PATCH 067/174] feat: add `mapsTo_fst_prod` and `mapsTo_snd_prod` (#17279) Also adjust the file-level docstring of `Data.Set.Prod` to match the content of the file --- Mathlib/Data/Set/Prod.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index acc80442bafa7..b1c42bc17e180 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -9,13 +9,15 @@ import Mathlib.Data.SProd /-! # Sets in product and pi types -This file defines the product of sets in `α × β` and in `Π i, α i` along with the diagonal of a -type. +This file proves basic properties of product of sets in `α × β` and in `Π i, α i`, and of the +diagonal of a type. ## Main declarations +This file contains basic results on the following notions, which are defined in `Set.Operations`. + * `Set.prod`: Binary product of sets. For `s : Set α`, `t : Set β`, we have - `s.prod t : Set (α × β)`. + `s.prod t : Set (α × β)`. Denoted by `s ×ˢ t`. * `Set.diagonal`: Diagonal of a type. `Set.diagonal α = {(x, x) | x : α}`. * `Set.offDiag`: Off-diagonal. `s ×ˢ s` without the diagonal. * `Set.pi`: Arbitrary product of sets. @@ -299,6 +301,9 @@ theorem fst_image_prod (s : Set β) {t : Set α} (ht : t.Nonempty) : Prod.fst '' let ⟨x, hx⟩ := ht ⟨(y, x), ⟨hy, hx⟩, rfl⟩ +lemma mapsTo_fst_prod {s : Set α} {t : Set β} : MapsTo Prod.fst (s ×ˢ t) s := + fun _ hx ↦ (mem_prod.1 hx).1 + theorem prod_subset_preimage_snd (s : Set α) (t : Set β) : s ×ˢ t ⊆ Prod.snd ⁻¹' t := inter_subset_right @@ -310,6 +315,9 @@ theorem snd_image_prod {s : Set α} (hs : s.Nonempty) (t : Set β) : Prod.snd '' let ⟨x, x_in⟩ := hs ⟨(x, y), ⟨x_in, y_in⟩, rfl⟩ +lemma mapsTo_snd_prod {s : Set α} {t : Set β} : MapsTo Prod.snd (s ×ˢ t) t := + fun _ hx ↦ (mem_prod.1 hx).2 + theorem prod_diff_prod : s ×ˢ t \ s₁ ×ˢ t₁ = s ×ˢ (t \ t₁) ∪ (s \ s₁) ×ˢ t := by ext x by_cases h₁ : x.1 ∈ s₁ <;> by_cases h₂ : x.2 ∈ t₁ <;> simp [*] From 445d05d7217497e72fc78f57cf2f7b2e784d2a39 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 17:40:33 +0000 Subject: [PATCH 068/174] feat: VC-dimension is monotone (#17283) From LeanCamCombi --- Mathlib/Combinatorics/SetFamily/Shatter.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/SetFamily/Shatter.lean b/Mathlib/Combinatorics/SetFamily/Shatter.lean index 6efd8f20ec42e..74f21f70fbc02 100644 --- a/Mathlib/Combinatorics/SetFamily/Shatter.lean +++ b/Mathlib/Combinatorics/SetFamily/Shatter.lean @@ -79,7 +79,7 @@ def shatterer (𝒜 : Finset (Finset α)) : Finset (Finset α) := (𝒜.biUnion simp_rw [mem_biUnion, mem_powerset] exact h.exists_superset -lemma shatterer_mono (h : 𝒜 ⊆ ℬ) : 𝒜.shatterer ⊆ ℬ.shatterer := +@[gcongr] lemma shatterer_mono (h : 𝒜 ⊆ ℬ) : 𝒜.shatterer ⊆ ℬ.shatterer := fun _ ↦ by simpa using Shatters.mono_left h lemma subset_shatterer (h : IsLowerSet (𝒜 : Set (Finset α))) : 𝒜 ⊆ 𝒜.shatterer := @@ -181,6 +181,8 @@ lemma shatterer_compress_subset_shatterer (a : α) (𝒜 : Finset (Finset α)) : /-- The Vapnik-Chervonenkis dimension of a set family is the maximal size of a set it shatters. -/ 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 /-- Down-compressing decreases the VC-dimension. -/ From 5024c93693c201b2cb4bb166bfea17cbda5e706e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Sep 2024 17:40:35 +0000 Subject: [PATCH 069/174] =?UTF-8?q?feat:=20`Disjoint=20G=E2=82=81.edgeFins?= =?UTF-8?q?et=20G=E2=82=82.edgeFinset=20=E2=86=94=20Disjoint=20G=E2=82=81?= =?UTF-8?q?=20G=E2=82=82`=20(#17286)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi --- Mathlib/Combinatorics/SimpleGraph/Finite.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Finite.lean b/Mathlib/Combinatorics/SimpleGraph/Finite.lean index 278ba02e0df80..bab7f40b821df 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Finite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Finite.lean @@ -92,6 +92,15 @@ theorem edgeFinset_inf [DecidableEq V] : (G₁ ⊓ G₂).edgeFinset = G₁.edgeF theorem edgeFinset_sdiff [DecidableEq V] : (G₁ \ G₂).edgeFinset = G₁.edgeFinset \ G₂.edgeFinset := by simp [edgeFinset] +lemma disjoint_edgeFinset : Disjoint G₁.edgeFinset G₂.edgeFinset ↔ Disjoint G₁ G₂ := by + simp_rw [← Finset.disjoint_coe, coe_edgeFinset, disjoint_edgeSet] + +lemma edgeFinset_eq_empty : G.edgeFinset = ∅ ↔ G = ⊥ := by + rw [← edgeFinset_bot, edgeFinset_inj] + +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 := Set.toFinset_card _ From d5d186a1c10fce0afec56557791b69d1e32546fa Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Mon, 30 Sep 2024 20:40:41 +0000 Subject: [PATCH 070/174] feat(Algebra/Group/Pointwise/Set/Basic): sInter and sUnion forms of pointwise operations (#17253) Provides `sInter` and `sUnion` versions of pointwise algebraic operations on sets. Needed in #17029 Co-authored-by: Christopher Hoskin --- .../Algebra/Group/Pointwise/Set/Basic.lean | 166 +++++++++++++----- Mathlib/Data/Set/Lattice.lean | 26 +++ 2 files changed, 147 insertions(+), 45 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 1ab1dcd9ea7a1..df2bf0c9e4874 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -181,10 +181,18 @@ theorem union_inv : (s ∪ t)⁻¹ = s⁻¹ ∪ t⁻¹ := theorem iInter_inv (s : ι → Set α) : (⋂ i, s i)⁻¹ = ⋂ i, (s i)⁻¹ := preimage_iInter +@[to_additive (attr := simp)] +theorem sInter_inv (S : Set (Set α)) : (⋂₀ S)⁻¹ = ⋂ s ∈ S, s⁻¹ := + preimage_sInter + @[to_additive (attr := simp)] theorem iUnion_inv (s : ι → Set α) : (⋃ i, s i)⁻¹ = ⋃ i, (s i)⁻¹ := preimage_iUnion +@[to_additive (attr := simp)] +theorem sUnion_inv (S : Set (Set α)) : (⋃₀ S)⁻¹ = ⋃ s ∈ S, s⁻¹ := + preimage_sUnion + @[to_additive (attr := simp)] theorem compl_inv : sᶜ⁻¹ = s⁻¹ᶜ := preimage_compl @@ -236,7 +244,7 @@ theorem inv_insert (a : α) (s : Set α) : (insert a s)⁻¹ = insert a⁻¹ s @[to_additive] theorem inv_range {ι : Sort*} {f : ι → α} : (range f)⁻¹ = range fun i => (f i)⁻¹ := by rw [← image_inv] - exact (range_comp _ _).symm + exact (range_comp ..).symm open MulOpposite @@ -374,47 +382,63 @@ theorem iUnion_mul_right_image : ⋃ a ∈ t, (· * a) '' s = s * t := @[to_additive] theorem iUnion_mul (s : ι → Set α) (t : Set α) : (⋃ i, s i) * t = ⋃ i, s i * t := - image2_iUnion_left _ _ _ + image2_iUnion_left .. @[to_additive] theorem mul_iUnion (s : Set α) (t : ι → Set α) : (s * ⋃ i, t i) = ⋃ i, s * t i := - image2_iUnion_right _ _ _ + image2_iUnion_right .. + +@[to_additive] +theorem sUnion_mul (S : Set (Set α)) (t : Set α) : ⋃₀ S * t = ⋃ s ∈ S, s * t := + image2_sUnion_left .. + +@[to_additive] +theorem mul_sUnion (s : Set α) (T : Set (Set α)) : s * ⋃₀ T = ⋃ t ∈ T, s * t := + image2_sUnion_right .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem iUnion₂_mul (s : ∀ i, κ i → Set α) (t : Set α) : (⋃ (i) (j), s i j) * t = ⋃ (i) (j), s i j * t := - image2_iUnion₂_left _ _ _ + image2_iUnion₂_left .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem mul_iUnion₂ (s : Set α) (t : ∀ i, κ i → Set α) : (s * ⋃ (i) (j), t i j) = ⋃ (i) (j), s * t i j := - image2_iUnion₂_right _ _ _ + image2_iUnion₂_right .. @[to_additive] theorem iInter_mul_subset (s : ι → Set α) (t : Set α) : (⋂ i, s i) * t ⊆ ⋂ i, s i * t := - Set.image2_iInter_subset_left _ _ _ + Set.image2_iInter_subset_left .. @[to_additive] theorem mul_iInter_subset (s : Set α) (t : ι → Set α) : (s * ⋂ i, t i) ⊆ ⋂ i, s * t i := - image2_iInter_subset_right _ _ _ + image2_iInter_subset_right .. + +@[to_additive] +lemma mul_sInter_subset (s : Set α) (T : Set (Set α)) : + s * ⋂₀ T ⊆ ⋂ t ∈ T, s * t := image2_sInter_right_subset s T (fun a b => a * b) + +@[to_additive] +lemma sInter_mul_subset (S : Set (Set α)) (t : Set α) : + ⋂₀ S * t ⊆ ⋂ s ∈ S, s * t := image2_sInter_left_subset S t (fun a b => a * b) /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem iInter₂_mul_subset (s : ∀ i, κ i → Set α) (t : Set α) : (⋂ (i) (j), s i j) * t ⊆ ⋂ (i) (j), s i j * t := - image2_iInter₂_subset_left _ _ _ + image2_iInter₂_subset_left .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem mul_iInter₂_subset (s : Set α) (t : ∀ i, κ i → Set α) : (s * ⋂ (i) (j), t i j) ⊆ ⋂ (i) (j), s * t i j := - image2_iInter₂_subset_right _ _ _ + image2_iInter₂_subset_right .. /-- The singleton operation as a `MulHom`. -/ @[to_additive "The singleton operation as an `AddHom`."] @@ -562,47 +586,63 @@ theorem iUnion_div_right_image : ⋃ a ∈ t, (· / a) '' s = s / t := @[to_additive] theorem iUnion_div (s : ι → Set α) (t : Set α) : (⋃ i, s i) / t = ⋃ i, s i / t := - image2_iUnion_left _ _ _ + image2_iUnion_left .. @[to_additive] theorem div_iUnion (s : Set α) (t : ι → Set α) : (s / ⋃ i, t i) = ⋃ i, s / t i := - image2_iUnion_right _ _ _ + image2_iUnion_right .. + +@[to_additive] +theorem sUnion_div (S : Set (Set α)) (t : Set α) : ⋃₀ S / t = ⋃ s ∈ S, s / t := + image2_sUnion_left .. + +@[to_additive] +theorem div_sUnion (s : Set α) (T : Set (Set α)) : s / ⋃₀ T = ⋃ t ∈ T, s / t := + image2_sUnion_right .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem iUnion₂_div (s : ∀ i, κ i → Set α) (t : Set α) : (⋃ (i) (j), s i j) / t = ⋃ (i) (j), s i j / t := - image2_iUnion₂_left _ _ _ + image2_iUnion₂_left .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem div_iUnion₂ (s : Set α) (t : ∀ i, κ i → Set α) : (s / ⋃ (i) (j), t i j) = ⋃ (i) (j), s / t i j := - image2_iUnion₂_right _ _ _ + image2_iUnion₂_right .. @[to_additive] theorem iInter_div_subset (s : ι → Set α) (t : Set α) : (⋂ i, s i) / t ⊆ ⋂ i, s i / t := - image2_iInter_subset_left _ _ _ + image2_iInter_subset_left .. @[to_additive] theorem div_iInter_subset (s : Set α) (t : ι → Set α) : (s / ⋂ i, t i) ⊆ ⋂ i, s / t i := - image2_iInter_subset_right _ _ _ + image2_iInter_subset_right .. + +@[to_additive] +theorem sInter_div_subset (S : Set (Set α)) (t : Set α) : ⋂₀ S / t ⊆ ⋂ s ∈ S, s / t := + image2_sInter_subset_left .. + +@[to_additive] +theorem div_sInter_subset (s : Set α) (T : Set (Set α)) : s / ⋂₀ T ⊆ ⋂ t ∈ T, s / t := + image2_sInter_subset_right .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem iInter₂_div_subset (s : ∀ i, κ i → Set α) (t : Set α) : (⋂ (i) (j), s i j) / t ⊆ ⋂ (i) (j), s i j / t := - image2_iInter₂_subset_left _ _ _ + image2_iInter₂_subset_left .. /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[to_additive] theorem div_iInter₂_subset (s : Set α) (t : ∀ i, κ i → Set α) : (s / ⋂ (i) (j), t i j) ⊆ ⋂ (i) (j), s / t i j := - image2_iInter₂_subset_right _ _ _ + image2_iInter₂_subset_right .. end Div @@ -691,35 +731,51 @@ lemma iUnion_smul_right_image : ⋃ a ∈ t, (· • a) '' s = s • t := iUnion @[to_additive] lemma iUnion_smul (s : ι → Set α) (t : Set β) : (⋃ i, s i) • t = ⋃ i, s i • t := - image2_iUnion_left _ _ _ + image2_iUnion_left .. @[to_additive] lemma smul_iUnion (s : Set α) (t : ι → Set β) : (s • ⋃ i, t i) = ⋃ i, s • t i := - image2_iUnion_right _ _ _ + image2_iUnion_right .. + +@[to_additive] +lemma sUnion_smul (S : Set (Set α)) (t : Set β) : ⋃₀ S • t = ⋃ s ∈ S, s • t := + image2_sUnion_left .. + +@[to_additive] +lemma smul_sUnion (s : Set α) (T : Set (Set β)) : s • ⋃₀ T = ⋃ t ∈ T, s • t := + image2_sUnion_right .. @[to_additive] lemma iUnion₂_smul (s : ∀ i, κ i → Set α) (t : Set β) : - (⋃ i, ⋃ j, s i j) • t = ⋃ i, ⋃ j, s i j • t := image2_iUnion₂_left _ _ _ + (⋃ i, ⋃ j, s i j) • t = ⋃ i, ⋃ j, s i j • t := image2_iUnion₂_left .. @[to_additive] lemma smul_iUnion₂ (s : Set α) (t : ∀ i, κ i → Set β) : - (s • ⋃ i, ⋃ j, t i j) = ⋃ i, ⋃ j, s • t i j := image2_iUnion₂_right _ _ _ + (s • ⋃ i, ⋃ j, t i j) = ⋃ i, ⋃ j, s • t i j := image2_iUnion₂_right .. @[to_additive] lemma iInter_smul_subset (s : ι → Set α) (t : Set β) : (⋂ i, s i) • t ⊆ ⋂ i, s i • t := - image2_iInter_subset_left _ _ _ + image2_iInter_subset_left .. @[to_additive] lemma smul_iInter_subset (s : Set α) (t : ι → Set β) : (s • ⋂ i, t i) ⊆ ⋂ i, s • t i := - image2_iInter_subset_right _ _ _ + image2_iInter_subset_right .. + +@[to_additive] +lemma sInter_smul_subset (S : Set (Set α)) (t : Set β) : ⋂₀ S • t ⊆ ⋂ s ∈ S, s • t := + image2_sInter_left_subset S t (fun a x => a • x) + +@[to_additive] +lemma smul_sInter_subset (s : Set α) (T : Set (Set β)) : s • ⋂₀ T ⊆ ⋂ t ∈ T, s • t := + image2_sInter_right_subset s T (fun a x => a • x) @[to_additive] lemma iInter₂_smul_subset (s : ∀ i, κ i → Set α) (t : Set β) : - (⋂ i, ⋂ j, s i j) • t ⊆ ⋂ i, ⋂ j, s i j • t := image2_iInter₂_subset_left _ _ _ + (⋂ i, ⋂ j, s i j) • t ⊆ ⋂ i, ⋂ j, s i j • t := image2_iInter₂_subset_left .. @[to_additive] lemma smul_iInter₂_subset (s : Set α) (t : ∀ i, κ i → Set β) : - (s • ⋂ i, ⋂ j, t i j) ⊆ ⋂ i, ⋂ j, s • t i j := image2_iInter₂_subset_right _ _ _ + (s • ⋂ i, ⋂ j, t i j) ⊆ ⋂ i, ⋂ j, s • t i j := image2_iInter₂_subset_right .. @[to_additive] lemma smul_set_subset_smul {s : Set α} : a ∈ s → a • t ⊆ s • t := image_subset_image2_right @@ -757,7 +813,7 @@ lemma smul_set_subset_iff : a • s ⊆ t ↔ ∀ ⦃b⦄, b ∈ s → a • b @[to_additive] lemma smul_set_union : a • (t₁ ∪ t₂) = a • t₁ ∪ a • t₂ := - image_union _ _ _ + image_union .. @[to_additive] lemma smul_set_insert (a : α) (b : β) (s : Set β) : a • insert b s = insert (a • b) (a • s) := @@ -765,7 +821,7 @@ lemma smul_set_insert (a : α) (b : β) (s : Set β) : a • insert b s = insert @[to_additive] lemma smul_set_inter_subset : a • (t₁ ∩ t₂) ⊆ a • t₁ ∩ a • t₂ := - image_inter_subset _ _ _ + image_inter_subset .. @[to_additive] lemma smul_set_iUnion (a : α) (s : ι → Set β) : a • ⋃ i, s i = ⋃ i, a • s i := @@ -773,15 +829,23 @@ lemma smul_set_iUnion (a : α) (s : ι → Set β) : a • ⋃ i, s i = ⋃ i, a @[to_additive] lemma smul_set_iUnion₂ (a : α) (s : ∀ i, κ i → Set β) : - a • ⋃ i, ⋃ j, s i j = ⋃ i, ⋃ j, a • s i j := image_iUnion₂ _ _ + a • ⋃ i, ⋃ j, s i j = ⋃ i, ⋃ j, a • s i j := image_iUnion₂ .. + +@[to_additive] +lemma smul_set_sUnion (a : α) (S : Set (Set β)) : a • ⋃₀ S = ⋃ s ∈ S, a • s := by + rw [sUnion_eq_biUnion, smul_set_iUnion₂] @[to_additive] lemma smul_set_iInter_subset (a : α) (t : ι → Set β) : a • ⋂ i, t i ⊆ ⋂ i, a • t i := - image_iInter_subset _ _ + image_iInter_subset .. + +@[to_additive] +lemma smul_set_sInter_subset (a : α) (S : Set (Set β)) : + a • ⋂₀ S ⊆ ⋂ s ∈ S, a • s := image_sInter_subset .. @[to_additive] lemma smul_set_iInter₂_subset (a : α) (t : ∀ i, κ i → Set β) : - a • ⋂ i, ⋂ j, t i j ⊆ ⋂ i, ⋂ j, a • t i j := image_iInter₂_subset _ _ + a • ⋂ i, ⋂ j, t i j ⊆ ⋂ i, ⋂ j, a • t i j := image_iInter₂_subset .. @[to_additive] lemma Nonempty.smul_set : s.Nonempty → (a • s).Nonempty := Nonempty.image _ @@ -797,7 +861,7 @@ lemma range_smul_range {ι κ : Type*} [SMul α β] (b : ι → α) (c : κ → @[to_additive] lemma smul_set_range [SMul α β] {ι : Sort*} (a : α) (f : ι → β) : a • range f = range fun i ↦ a • f i := - (range_comp _ _).symm + (range_comp ..).symm @[to_additive] lemma range_smul [SMul α β] {ι : Sort*} (a : α) (f : ι → β) : range (fun i ↦ a • f i) = a • range f := (smul_set_range ..).symm @@ -863,28 +927,40 @@ lemma iUnion_vsub_left_image : ⋃ a ∈ s, (a -ᵥ ·) '' t = s -ᵥ t := iUnio lemma iUnion_vsub_right_image : ⋃ a ∈ t, (· -ᵥ a) '' s = s -ᵥ t := iUnion_image_right _ lemma iUnion_vsub (s : ι → Set β) (t : Set β) : (⋃ i, s i) -ᵥ t = ⋃ i, s i -ᵥ t := - image2_iUnion_left _ _ _ + image2_iUnion_left .. lemma vsub_iUnion (s : Set β) (t : ι → Set β) : (s -ᵥ ⋃ i, t i) = ⋃ i, s -ᵥ t i := - image2_iUnion_right _ _ _ + image2_iUnion_right .. + +lemma sUnion_vsub (S : Set (Set β)) (t : Set β) : ⋃₀ S -ᵥ t = ⋃ s ∈ S, s -ᵥ t := + image2_sUnion_left .. + +lemma vsub_sUnion (s : Set β) (T : Set (Set β)) : s -ᵥ ⋃₀ T = ⋃ t ∈ T, s -ᵥ t := + image2_sUnion_right .. lemma iUnion₂_vsub (s : ∀ i, κ i → Set β) (t : Set β) : - (⋃ i, ⋃ j, s i j) -ᵥ t = ⋃ i, ⋃ j, s i j -ᵥ t := image2_iUnion₂_left _ _ _ + (⋃ i, ⋃ j, s i j) -ᵥ t = ⋃ i, ⋃ j, s i j -ᵥ t := image2_iUnion₂_left .. lemma vsub_iUnion₂ (s : Set β) (t : ∀ i, κ i → Set β) : - (s -ᵥ ⋃ i, ⋃ j, t i j) = ⋃ i, ⋃ j, s -ᵥ t i j := image2_iUnion₂_right _ _ _ + (s -ᵥ ⋃ i, ⋃ j, t i j) = ⋃ i, ⋃ j, s -ᵥ t i j := image2_iUnion₂_right .. lemma iInter_vsub_subset (s : ι → Set β) (t : Set β) : (⋂ i, s i) -ᵥ t ⊆ ⋂ i, s i -ᵥ t := - image2_iInter_subset_left _ _ _ + image2_iInter_subset_left .. lemma vsub_iInter_subset (s : Set β) (t : ι → Set β) : (s -ᵥ ⋂ i, t i) ⊆ ⋂ i, s -ᵥ t i := - image2_iInter_subset_right _ _ _ + image2_iInter_subset_right .. + +lemma sInter_vsub_subset (S : Set (Set β)) (t : Set β) : ⋂₀ S -ᵥ t ⊆ ⋂ s ∈ S, s -ᵥ t := + image2_sInter_subset_left .. + +lemma vsub_sInter_subset (s : Set β) (T : Set (Set β)) : s -ᵥ ⋂₀ T ⊆ ⋂ t ∈ T, s -ᵥ t := + image2_sInter_subset_right .. lemma iInter₂_vsub_subset (s : ∀ i, κ i → Set β) (t : Set β) : - (⋂ i, ⋂ j, s i j) -ᵥ t ⊆ ⋂ i, ⋂ j, s i j -ᵥ t := image2_iInter₂_subset_left _ _ _ + (⋂ i, ⋂ j, s i j) -ᵥ t ⊆ ⋂ i, ⋂ j, s i j -ᵥ t := image2_iInter₂_subset_left .. lemma vsub_iInter₂_subset (s : Set β) (t : ∀ i, κ i → Set β) : - s -ᵥ ⋂ i, ⋂ j, t i j ⊆ ⋂ i, ⋂ j, s -ᵥ t i j := image2_iInter₂_subset_right _ _ _ + s -ᵥ ⋂ i, ⋂ j, t i j ⊆ ⋂ i, ⋂ j, s -ᵥ t i j := image2_iInter₂_subset_right .. end VSub @@ -1195,12 +1271,12 @@ theorem preimage_mul_right_one' : (· * b⁻¹) ⁻¹' 1 = {b} := by simp @[to_additive (attr := simp)] theorem mul_univ (hs : s.Nonempty) : s * (univ : Set α) = univ := let ⟨a, ha⟩ := hs - eq_univ_of_forall fun b => ⟨a, ha, a⁻¹ * b, trivial, mul_inv_cancel_left _ _⟩ + eq_univ_of_forall fun b => ⟨a, ha, a⁻¹ * b, trivial, mul_inv_cancel_left ..⟩ @[to_additive (attr := simp)] theorem univ_mul (ht : t.Nonempty) : (univ : Set α) * t = univ := let ⟨a, ha⟩ := ht - eq_univ_of_forall fun b => ⟨b * a⁻¹, trivial, a, ha, inv_mul_cancel_right _ _⟩ + eq_univ_of_forall fun b => ⟨b * a⁻¹, trivial, a, ha, inv_mul_cancel_right ..⟩ end Group @@ -1217,12 +1293,12 @@ lemma mul_subset_range {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) rintro _ ⟨a, ha, b, hb, rfl⟩ obtain ⟨a, rfl⟩ := hs ha obtain ⟨b, rfl⟩ := ht hb - exact ⟨a * b, map_mul _ _ _⟩ + exact ⟨a * b, map_mul ..⟩ @[to_additive] theorem preimage_mul_preimage_subset {s t : Set β} : m ⁻¹' s * m ⁻¹' t ⊆ m ⁻¹' (s * t) := by rintro _ ⟨_, _, _, _, rfl⟩ - exact ⟨_, ‹_›, _, ‹_›, (map_mul m _ _).symm⟩ + exact ⟨_, ‹_›, _, ‹_›, (map_mul m ..).symm⟩ @[to_additive] lemma preimage_mul (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : @@ -1246,12 +1322,12 @@ lemma div_subset_range {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) rintro _ ⟨a, ha, b, hb, rfl⟩ obtain ⟨a, rfl⟩ := hs ha obtain ⟨b, rfl⟩ := ht hb - exact ⟨a / b, map_div _ _ _⟩ + exact ⟨a / b, map_div ..⟩ @[to_additive] theorem preimage_div_preimage_subset {s t : Set β} : m ⁻¹' s / m ⁻¹' t ⊆ m ⁻¹' (s / t) := by rintro _ ⟨_, _, _, _, rfl⟩ - exact ⟨_, ‹_›, _, ‹_›, (map_div m _ _).symm⟩ + exact ⟨_, ‹_›, _, ‹_›, (map_div m ..).symm⟩ @[to_additive] lemma preimage_div (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index fbff041791382..b6f29fa406fdf 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -1243,6 +1243,14 @@ theorem image_sInter_subset (S : Set (Set α)) (f : α → β) : f '' ⋂₀ S rw [sInter_eq_biInter] apply image_iInter₂_subset +theorem image2_sInter_right_subset (t : Set α) (S : Set (Set β)) (f : α → β → γ) : + image2 f t (⋂₀ S) ⊆ ⋂ s ∈ S, image2 f t s := by + aesop + +theorem image2_sInter_left_subset (S : Set (Set α)) (t : Set β) (f : α → β → γ) : + image2 f (⋂₀ S) t ⊆ ⋂ s ∈ S, image2 f s t := by + aesop + /-! ### `restrictPreimage` -/ @@ -1575,6 +1583,14 @@ theorem image2_iUnion_right (s : Set α) (t : ι → Set β) : image2 f s (⋃ i, t i) = ⋃ i, image2 f s (t i) := by simp only [← image_prod, prod_iUnion, image_iUnion] +theorem image2_sUnion_left (S : Set (Set α)) (t : Set β) : + image2 f (⋃₀ S) t = ⋃ s ∈ S, image2 f s t := by + aesop + +theorem image2_sUnion_right (s : Set α) (T : Set (Set β)) : + image2 f s (⋃₀ T) = ⋃ t ∈ T, image2 f s t := by + aesop + /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ theorem image2_iUnion₂_left (s : ∀ i, κ i → Set α) (t : Set β) : @@ -1610,6 +1626,16 @@ theorem image2_iInter₂_subset_right (s : Set α) (t : ∀ i, κ i → Set β) simp_rw [image2_subset_iff, mem_iInter] exact fun x hx y hy i j => mem_image2_of_mem hx (hy _ _) +theorem image2_sInter_subset_left (S : Set (Set α)) (t : Set β) : + image2 f (⋂₀ S) t ⊆ ⋂ s ∈ S, image2 f s t := by + rw [sInter_eq_biInter] + exact image2_iInter₂_subset_left .. + +theorem image2_sInter_subset_right (s : Set α) (T : Set (Set β)) : + image2 f s (⋂₀ T) ⊆ ⋂ t ∈ T, image2 f s t := by + rw [sInter_eq_biInter] + exact image2_iInter₂_subset_right .. + theorem prod_eq_biUnion_left : s ×ˢ t = ⋃ a ∈ s, (fun b => (a, b)) '' t := by rw [iUnion_image_left, image2_mk_eq_prod] From 37814caf0b1b93a00743c1dd7af97ceb6b092b40 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 30 Sep 2024 21:32:25 +0000 Subject: [PATCH 071/174] feat(SetTheory/ZFC/Basic): make `sInter` computable (#17248) Moved from #15624 with permission from the original author. Co-authored by: RustyYato --- Mathlib/SetTheory/ZFC/Basic.lean | 20 +++++++++++++------- 1 file changed, 13 insertions(+), 7 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index bede4eaab6232..137369f0640e1 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -834,6 +834,10 @@ theorem mem_sep {p : ZFSet.{u} → Prop} {x y : ZFSet.{u}} : Quotient.inductionOn₂ x y fun _ _ => PSet.mem_sep (p := p ∘ mk) fun _ _ h => (Quotient.sound h).subst +@[simp] +theorem sep_empty (p : ZFSet → Prop) : (∅ : ZFSet).sep p = ∅ := + (eq_empty _).mpr fun _ h ↦ not_mem_empty _ (mem_sep.mp h).1 + @[simp] theorem toSet_sep (a : ZFSet) (p : ZFSet → Prop) : (ZFSet.sep p a).toSet = { x ∈ a.toSet | p x } := by @@ -886,9 +890,8 @@ def sUnion : ZFSet → ZFSet := prefix:110 "⋃₀ " => ZFSet.sUnion /-- The intersection operator, the collection of elements in all of the elements of a ZFC set. We -special-case `⋂₀ ∅ = ∅`. -/ -noncomputable def sInter (x : ZFSet) : ZFSet := by - classical exact if h : x.Nonempty then ZFSet.sep (fun y => ∀ z ∈ x, y ∈ z) h.some else ∅ +define `⋂₀ ∅ = ∅`. -/ +def sInter (x : ZFSet) : ZFSet := (⋃₀ x).sep (fun y => ∀ z ∈ x, y ∈ z) @[inherit_doc] prefix:110 "⋂₀ " => ZFSet.sInter @@ -899,9 +902,12 @@ theorem mem_sUnion {x y : ZFSet.{u}} : y ∈ ⋃₀ x ↔ ∃ z ∈ x, y ∈ z : ⟨fun ⟨z, h⟩ => ⟨⟦z⟧, h⟩, fun ⟨z, h⟩ => Quotient.inductionOn z (fun z h => ⟨z, h⟩) h⟩ theorem mem_sInter {x y : ZFSet} (h : x.Nonempty) : y ∈ ⋂₀ x ↔ ∀ z ∈ x, y ∈ z := by - rw [sInter, dif_pos h] - simp only [mem_toSet, mem_sep, and_iff_right_iff_imp] - exact fun H => H _ h.some_mem + unfold sInter + simp only [and_iff_right_iff_imp, mem_sep] + intro mem + apply mem_sUnion.mpr + replace ⟨s, h⟩ := h + exact ⟨_, h, mem _ h⟩ @[simp] theorem sUnion_empty : ⋃₀ (∅ : ZFSet.{u}) = ∅ := by @@ -909,7 +915,7 @@ theorem sUnion_empty : ⋃₀ (∅ : ZFSet.{u}) = ∅ := by simp @[simp] -theorem sInter_empty : ⋂₀ (∅ : ZFSet) = ∅ := dif_neg <| by simp +theorem sInter_empty : ⋂₀ (∅ : ZFSet) = ∅ := by simp [sInter] theorem mem_of_mem_sInter {x y z : ZFSet} (hy : y ∈ ⋂₀ x) (hz : z ∈ x) : y ∈ z := by rcases eq_empty_or_nonempty x with (rfl | hx) From 0da2361bd278b7158de3a93c1e39b29b63aae58c Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 1 Oct 2024 05:27:32 +0000 Subject: [PATCH 072/174] chore: tidy various files (#17095) --- Mathlib/Algebra/Algebra/Basic.lean | 6 +-- Mathlib/Algebra/Group/Subgroup/Basic.lean | 2 +- Mathlib/Algebra/Lie/Subalgebra.lean | 2 +- Mathlib/Algebra/Module/Equiv/Basic.lean | 2 +- .../SpecialFunctions/Gamma/Basic.lean | 3 +- .../Trigonometric/Bounds.lean | 9 ++-- Mathlib/FieldTheory/Finiteness.lean | 2 +- Mathlib/FieldTheory/Galois/Basic.lean | 2 +- .../GroupTheory/MonoidLocalization/Basic.lean | 37 +++++++------- Mathlib/LinearAlgebra/LinearIndependent.lean | 3 +- Mathlib/Logic/Function/FiberPartition.lean | 2 +- .../Function/LocallyIntegrable.lean | 6 +-- .../Measure/SeparableMeasure.lean | 49 +++++++++---------- .../NumberField/CanonicalEmbedding/Basic.lean | 2 +- .../CanonicalEmbedding/FundamentalCone.lean | 8 +-- Mathlib/Probability/Distributions/Gamma.lean | 9 ++-- .../RingTheory/MvPowerSeries/LexOrder.lean | 7 +-- .../RingTheory/UniqueFactorizationDomain.lean | 7 ++- Mathlib/Topology/FiberPartition.lean | 6 +-- 19 files changed, 73 insertions(+), 91 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 29d91b4ab7439..a7a16de103ee9 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -185,8 +185,7 @@ theorem End_algebraMap_isUnit_inv_apply_eq_iff {x : R} mpr := fun H => H.symm ▸ by apply_fun ⇑h.unit.val using ((Module.End_isUnit_iff _).mp h).injective - erw [End_isUnit_apply_inv_apply_of_isUnit] - rfl } + simpa using End_isUnit_apply_inv_apply_of_isUnit h (x • m') } theorem End_algebraMap_isUnit_inv_apply_eq_iff' {x : R} (h : IsUnit (algebraMap R (Module.End S M) x)) (m m' : M) : @@ -195,8 +194,7 @@ theorem End_algebraMap_isUnit_inv_apply_eq_iff' {x : R} mpr := fun H => H.symm ▸ by apply_fun (↑h.unit : M → M) using ((Module.End_isUnit_iff _).mp h).injective - erw [End_isUnit_apply_inv_apply_of_isUnit] - rfl } + simpa using End_isUnit_apply_inv_apply_of_isUnit h (x • m') |>.symm } end diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 0d13d5ae72a71..2df0f53da5529 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -69,7 +69,7 @@ Definitions in the file: * `MonoidHom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` such that `f x = 1` -* `MonoidHom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that +* `MonoidHom.eqLocus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that `f x = g x` form a subgroup of `G` ## Implementation notes diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 649208ace2362..a1d215735203f 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -224,7 +224,7 @@ variable [Module R M] `L`, we may regard `M` as a Lie module of `L'` by restriction. -/ instance lieModule [LieModule R L M] : LieModule R L' M where smul_lie t x m := by - rw [coe_bracket_of_module]; erw [smul_lie]; simp only [coe_bracket_of_module] + rw [coe_bracket_of_module, Submodule.coe_smul_of_tower, smul_lie, coe_bracket_of_module] lie_smul t x m := by simp only [coe_bracket_of_module, lie_smul] /-- An `L`-equivariant map of Lie modules `M → N` is `L'`-equivariant for any Lie subalgebra diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 3d210451cac46..502f5ddd8ebdf 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -122,7 +122,7 @@ protected theorem smul_def (f : M ≃ₗ[R] M) (a : M) : f • a = f a := /-- `LinearEquiv.applyDistribMulAction` is faithful. -/ instance apply_faithfulSMul : FaithfulSMul (M ≃ₗ[R] M) M := - ⟨@fun _ _ ↦ LinearEquiv.ext⟩ + ⟨LinearEquiv.ext⟩ instance apply_smulCommClass [SMul S R] [SMul S M] [IsScalarTower S R M] : SMulCommClass S (M ≃ₗ[R] M) M where diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index c12393d6b86df..86b3f777ce844 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -553,8 +553,7 @@ lemma integral_rpow_mul_exp_neg_mul_Ioi {a r : ℝ} (ha : 0 < a) (hr : 0 < r) : convert integral_cpow_mul_exp_neg_mul_Ioi (by rwa [ofReal_re] : 0 < (a : ℂ).re) hr refine _root_.integral_ofReal.symm.trans <| setIntegral_congr measurableSet_Ioi (fun t ht ↦ ?_) norm_cast - rw [← ofReal_cpow (le_of_lt ht), RCLike.ofReal_mul] - rfl + simp_rw [← ofReal_cpow ht.le, RCLike.ofReal_mul, coe_algebraMap] open Lean.Meta Qq Mathlib.Meta.Positivity in /-- The `positivity` extension which identifies expressions of the form `Gamma a`. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index e0595d589e11d..e8e17232197f1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -173,11 +173,9 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by let U := Ico 0 (π / 2) have intU : interior U = Ioo 0 (π / 2) := interior_Ico have half_pi_pos : 0 < π / 2 := div_pos pi_pos two_pos - have cos_pos : ∀ {y : ℝ}, y ∈ U → 0 < cos y := by - intro y hy + have cos_pos {y : ℝ} (hy : y ∈ U) : 0 < cos y := by exact cos_pos_of_mem_Ioo (Ico_subset_Ioo_left (neg_lt_zero.mpr half_pi_pos) hy) - have sin_pos : ∀ {y : ℝ}, y ∈ interior U → 0 < sin y := by - intro y hy + have sin_pos {y : ℝ} (hy : y ∈ interior U) : 0 < sin y := by rw [intU] at hy exact sin_pos_of_mem_Ioo (Ioo_subset_Ioo_right (div_le_self pi_pos.le one_le_two) hy) have tan_cts_U : ContinuousOn tan U := by @@ -186,8 +184,7 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by simp only [mem_setOf_eq] exact (cos_pos hz).ne' have tan_minus_id_cts : ContinuousOn (fun y : ℝ => tan y - y) U := tan_cts_U.sub continuousOn_id - have deriv_pos : ∀ y : ℝ, y ∈ interior U → 0 < deriv (fun y' : ℝ => tan y' - y') y := by - intro y hy + have deriv_pos (y : ℝ) (hy : y ∈ interior U) : 0 < deriv (fun y' : ℝ => tan y' - y') y := by have := cos_pos (interior_subset hy) simp only [deriv_tan_sub_id y this.ne', one_div, gt_iff_lt, sub_pos] norm_cast diff --git a/Mathlib/FieldTheory/Finiteness.lean b/Mathlib/FieldTheory/Finiteness.lean index 6c2e5621a9272..d6cc69e04c7ae 100644 --- a/Mathlib/FieldTheory/Finiteness.lean +++ b/Mathlib/FieldTheory/Finiteness.lean @@ -72,7 +72,7 @@ theorem coeSort_finsetBasisIndex [IsNoetherian K V] : /-- In a noetherian module over a division ring, there exists a finite basis. This is indexed by the `Finset` `IsNoetherian.finsetBasisIndex`. This is in contrast to the result `finite_basis_index (Basis.ofVectorSpace K V)`, -which provides a set and a `Set.finite`. +which provides a set and a `Set.Finite`. -/ noncomputable def finsetBasis [IsNoetherian K V] : Basis (finsetBasisIndex K V) K V := (Basis.ofVectorSpace K V).reindex (by rw [coeSort_finsetBasisIndex]) diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 868454440fa9b..56ce629fc2bea 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -24,7 +24,7 @@ In this file we define Galois extensions as extensions which are both separable - `IntermediateField.fixingSubgroup_fixedField` : If `E/F` is finite dimensional (but not necessarily Galois) then `fixingSubgroup (fixedField H) = H` -- `IntermediateField.fixedField_fixingSubgroup`: If `E/F` is finite dimensional and Galois +- `IsGalois.fixedField_fixingSubgroup`: If `E/F` is finite dimensional and Galois then `fixedField (fixingSubgroup K) = K` Together, these two results prove the Galois correspondence. diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 988928ee6d456..b5db348da57a7 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -548,7 +548,7 @@ theorem mk'_spec' (x) (y : S) : f.toMap y * f.mk' x y = f.toMap x := by rw [mul_ @[to_additive] theorem eq_mk'_iff_mul_eq {x} {y : S} {z} : z = f.mk' x y ↔ z * f.toMap y = f.toMap x := - ⟨fun H ↦ by rw [H, mk'_spec], fun H ↦ by erw [mul_inv_right, H]⟩ + ⟨fun H ↦ by rw [H, mk'_spec], fun H ↦ by rw [mk', mul_inv_right, H]⟩ @[to_additive] theorem mk'_eq_iff_eq_mul {x} {y : S} {z} : f.mk' x y = z ↔ f.toMap x = z * f.toMap y := by @@ -773,9 +773,7 @@ theorem lift_comp : (f.lift hg).comp f.toMap = g := by ext; exact f.lift_eq hg _ @[to_additive (attr := simp)] theorem lift_of_comp (j : N →* P) : f.lift (f.isUnit_comp j) = j := by ext - rw [lift_spec] - show j _ = j _ * _ - erw [← j.map_mul, sec_spec'] + simp_rw [lift_spec, MonoidHom.comp_apply, ← j.map_mul, sec_spec'] @[to_additive] theorem epic_of_localizationMap {j k : N →* P} (h : ∀ a, j.comp f.toMap a = k.comp f.toMap a) : @@ -834,8 +832,8 @@ theorem lift_surjective_iff : obtain ⟨z, hz⟩ := H v obtain ⟨x, hx⟩ := f.surj z use x - rw [← hz, f.eq_mk'_iff_mul_eq.2 hx, lift_mk', mul_assoc, mul_comm _ (g ↑x.2)] - erw [IsUnit.mul_liftRight_inv (g.restrict S) hg, mul_one] + rw [← hz, f.eq_mk'_iff_mul_eq.2 hx, lift_mk', mul_assoc, mul_comm _ (g ↑x.2), + ← MonoidHom.restrict_apply, IsUnit.mul_liftRight_inv (g.restrict S) hg, mul_one] · intro H v obtain ⟨x, hx⟩ := H v use f.mk' x.1 x.2 @@ -1131,9 +1129,9 @@ of `AddCommMonoid`s, `k ∘ f` is a Localization map for `M` at `S`."] def ofMulEquivOfLocalizations (k : N ≃* P) : LocalizationMap S P := (k.toMonoidHom.comp f.toMap).toLocalizationMap (fun y ↦ isUnit_comp f k.toMonoidHom y) (fun v ↦ - let ⟨z, hz⟩ := k.toEquiv.surjective v + let ⟨z, hz⟩ := k.surjective v let ⟨x, hx⟩ := f.surj z - ⟨x, show v * k _ = k _ by rw [← hx, map_mul, ← hz]; rfl⟩) + ⟨x, show v * k _ = k _ by rw [← hx, map_mul, ← hz]⟩) fun x y ↦ (k.apply_eq_iff_eq.trans f.eq_iff_exists).1 @[to_additive (attr := simp)] @@ -1203,18 +1201,17 @@ def ofMulEquivOfDom {k : P ≃* M} (H : T.map k.toMonoidHom = S) : LocalizationM ⟨z, hz⟩) (fun z ↦ let ⟨x, hx⟩ := f.surj z - let ⟨v, hv⟩ := k.toEquiv.surjective x.1 - let ⟨w, hw⟩ := k.toEquiv.surjective x.2 - ⟨(v, ⟨w, H' ▸ show k w ∈ S from hw.symm ▸ x.2.2⟩), - show z * f.toMap (k.toEquiv w) = f.toMap (k.toEquiv v) by erw [hv, hw, hx]⟩) - fun x y ↦ - show f.toMap _ = f.toMap _ → _ by - erw [f.eq_iff_exists] - exact - fun ⟨c, hc⟩ ↦ - let ⟨d, hd⟩ := k.toEquiv.surjective c - ⟨⟨d, H' ▸ show k d ∈ S from hd.symm ▸ c.2⟩, by - erw [← hd, ← map_mul k, ← map_mul k] at hc; exact k.toEquiv.injective hc⟩ + let ⟨v, hv⟩ := k.surjective x.1 + let ⟨w, hw⟩ := k.surjective x.2 + ⟨(v, ⟨w, H' ▸ show k w ∈ S from hw.symm ▸ x.2.2⟩), by + simp_rw [MonoidHom.comp_apply, MulEquiv.toMonoidHom_eq_coe, MonoidHom.coe_coe, hv, hw, hx]⟩) + fun x y ↦ by + rw [MonoidHom.comp_apply, MonoidHom.comp_apply, MulEquiv.toMonoidHom_eq_coe, + MonoidHom.coe_coe, f.eq_iff_exists] + rintro ⟨c, hc⟩ + let ⟨d, hd⟩ := k.surjective c + refine ⟨⟨d, H' ▸ show k d ∈ S from hd.symm ▸ c.2⟩, ?_⟩ + rw [← hd, ← map_mul k, ← map_mul k] at hc; exact k.injective hc @[to_additive (attr := simp)] theorem ofMulEquivOfDom_apply {k : P ≃* M} (H : T.map k.toMonoidHom = S) (x) : diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index bba76de12bd95..ed68d3b7b7e44 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -591,8 +591,7 @@ theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) exact (hgs i hi).symm ▸ zero_smul _ _ · rw [← hsum, Finset.sum_congr rfl _] intros - erw [Pi.smul_apply, smul_assoc] - rfl + rw [Pi.smul_apply', smul_assoc, Units.smul_def] lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := by diff --git a/Mathlib/Logic/Function/FiberPartition.lean b/Mathlib/Logic/Function/FiberPartition.lean index e82895e5b975f..e1463b5d0e46b 100644 --- a/Mathlib/Logic/Function/FiberPartition.lean +++ b/Mathlib/Logic/Function/FiberPartition.lean @@ -62,7 +62,7 @@ lemma fiber_nonempty (f : Y → Z) (a : Fiber f) : Set.Nonempty a.val := by rw [mem_iff_eq_image, ← map_preimage_eq_image] lemma map_preimage_eq_image_map {W : Type*} (f : Y → Z) (g : Z → W) (a : Fiber (g ∘ f)) : - g (f a.preimage) = a.image := by rw [← map_preimage_eq_image]; rfl + g (f a.preimage) = a.image := by rw [← map_preimage_eq_image, comp_apply] lemma image_eq_image_mk (f : Y → Z) (g : X → Y) (a : Fiber (f ∘ g)) : a.image = (Fiber.mk f (g (a.preimage _))).image := by diff --git a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean index db051f36a5bc3..0d8110dc76977 100644 --- a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean @@ -127,10 +127,10 @@ theorem LocallyIntegrableOn.aestronglyMeasurable [SecondCountableTopology X] rw [this, aestronglyMeasurable_iUnion_iff] exact fun i : ℕ => (hu i).aestronglyMeasurable -/-- If `s` is either open, or closed, then `f` is locally integrable on `s` iff it is integrable on -every compact subset contained in `s`. -/ +/-- If `s` is locally closed (e.g. open or closed), then `f` is locally integrable on `s` iff it is +integrable on every compact subset contained in `s`. -/ theorem locallyIntegrableOn_iff [LocallyCompactSpace X] (hs : IsLocallyClosed s) : - LocallyIntegrableOn f s μ ↔ ∀ (k : Set X), k ⊆ s → (IsCompact k → IntegrableOn f k μ) := by + LocallyIntegrableOn f s μ ↔ ∀ (k : Set X), k ⊆ s → IsCompact k → IntegrableOn f k μ := by refine ⟨fun hf k hk ↦ hf.integrableOn_compact_subset hk, fun hf x hx ↦ ?_⟩ rcases hs with ⟨U, Z, hU, hZ, rfl⟩ rcases exists_compact_subset hU hx.1 with ⟨K, hK, hxK, hKU⟩ diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index a475581af381f..55e3abb732208 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -32,7 +32,7 @@ of separability in the metric space made by constant indicators equipped with th ## Main definitions -* `MeasureTheory.Measure.μ.MeasureDense 𝒜`: `𝒜` is a measure-dense family if it only contains +* `MeasureTheory.Measure.MeasureDense μ 𝒜`: `𝒜` is a measure-dense family if it only contains measurable sets and if the following condition is satisfied: if `s` is measurable with finite measure, then for any `ε > 0` there exists `t ∈ 𝒜` such that `μ (s ∆ t) < ε `. * `MeasureTheory.IsSeparable`: A measure is separable if there exists a countable and @@ -87,8 +87,7 @@ structure Measure.MeasureDense (μ : Measure X) (𝒜 : Set (Set X)) : Prop := /-- Each set has to be measurable. -/ measurable : ∀ s ∈ 𝒜, MeasurableSet s /-- Any measurable set can be approximated by sets in the family. -/ - approx : ∀ s, MeasurableSet s → μ s ≠ ∞ → ∀ (ε : ℝ), - 0 < ε → ∃ t ∈ 𝒜, μ (s ∆ t) < ENNReal.ofReal ε + approx : ∀ s, MeasurableSet s → μ s ≠ ∞ → ∀ ε : ℝ, 0 < ε → ∃ t ∈ 𝒜, μ (s ∆ t) < ENNReal.ofReal ε theorem Measure.MeasureDense.nonempty (h𝒜 : μ.MeasureDense 𝒜) : 𝒜.Nonempty := by rcases h𝒜.approx ∅ MeasurableSet.empty (by simp) 1 (by norm_num) with ⟨t, ht, -⟩ @@ -103,8 +102,8 @@ theorem Measure.MeasureDense.nonempty' (h𝒜 : μ.MeasureDense 𝒜) : /-- The set of measurable sets is measure-dense. -/ theorem measureDense_measurableSet : μ.MeasureDense {s | MeasurableSet s} where - measurable := fun _ h ↦ h - approx := fun s hs _ ε ε_pos ↦ ⟨s, hs, by simpa⟩ + measurable _ h := h + approx s hs _ ε ε_pos := ⟨s, hs, by simpa⟩ /-- If a family of sets `𝒜` is measure-dense in `X`, then any measurable set with finite measure can be approximated by sets in `𝒜` with finite measure. -/ @@ -155,9 +154,8 @@ theorem Measure.MeasureDense.indicatorConstLp_subset_closure (h𝒜 : μ.Measure with finite measure. -/ theorem Measure.MeasureDense.fin_meas (h𝒜 : μ.MeasureDense 𝒜) : μ.MeasureDense {s | s ∈ 𝒜 ∧ μ s ≠ ∞} where - measurable := fun s h ↦ h𝒜.measurable s h.1 - approx := by - intro s ms hμs ε ε_pos + measurable s h := h𝒜.measurable s h.1 + approx s ms hμs ε ε_pos := by rcases Measure.MeasureDense.fin_meas_approx h𝒜 ms hμs ε ε_pos with ⟨t, t_mem, hμt, hμst⟩ exact ⟨t, ⟨t_mem, hμt⟩, hμst⟩ @@ -165,12 +163,11 @@ theorem Measure.MeasureDense.fin_meas (h𝒜 : μ.MeasureDense 𝒜) : this algebra of sets is measure-dense. -/ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_finite [IsFiniteMeasure μ] (h𝒜 : IsSetAlgebra 𝒜) (hgen : m = MeasurableSpace.generateFrom 𝒜) : μ.MeasureDense 𝒜 where - measurable := fun s hs ↦ hgen ▸ measurableSet_generateFrom hs - approx := by + measurable s hs := hgen ▸ measurableSet_generateFrom hs + approx s ms := by -- We want to show that any measurable set can be approximated by sets in `𝒜`. To do so, it is -- enough to show that such sets constitute a `σ`-algebra containing `𝒜`. This is contained in -- the theorem `generateFrom_induction`. - intro s ms have : MeasurableSet s ∧ ∀ (ε : ℝ), 0 < ε → ∃ t ∈ 𝒜, (μ (s ∆ t)).toReal < ε := by apply generateFrom_induction (p := fun s ↦ MeasurableSet s ∧ ∀ (ε : ℝ), 0 < ε → ∃ t ∈ 𝒜, (μ (s ∆ t)).toReal < ε) @@ -220,14 +217,14 @@ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_finite [IsFiniteMeasur apply _root_.add_lt_add · rw [measure_diff (h_fin := measure_ne_top _ _), toReal_sub_of_le (ha := measure_ne_top _ _)] - apply lt_of_le_of_lt (sub_le_dist ..) - simp only [Finset.mem_range, Nat.lt_add_one_iff] - exact (dist_comm (α := ℝ) .. ▸ hN N (le_refl N)) - exact (measure_mono <| iUnion_subset <| - fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i)) - exact iUnion_subset <| fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i) - exact MeasurableSet.biUnion (countable_coe_iff.1 inferInstance) - (fun n _ ↦ (hf n).1.nullMeasurableSet) + · apply lt_of_le_of_lt (sub_le_dist ..) + simp only [Finset.mem_range, Nat.lt_add_one_iff] + exact (dist_comm (α := ℝ) .. ▸ hN N (le_refl N)) + · exact measure_mono <| iUnion_subset <| + fun i ↦ iUnion_subset fun _ ↦ subset_iUnion f i + · exact iUnion_subset <| fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i) + · exact MeasurableSet.biUnion (countable_coe_iff.1 inferInstance) + (fun n _ ↦ (hf n).1.nullMeasurableSet) · calc (μ ((⋃ n ∈ (Finset.range (N + 1)), f n) ∆ (⋃ n ∈ (Finset.range (N + 1)), g ↑n))).toReal @@ -257,18 +254,17 @@ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_sigmaFinite (h𝒜 : I (S : μ.FiniteSpanningSetsIn 𝒜) (hgen : m = MeasurableSpace.generateFrom 𝒜) : μ.MeasureDense 𝒜 where measurable s hs := hgen ▸ measurableSet_generateFrom hs - approx := by + approx s ms hμs ε ε_pos := by -- We use partial unions of (Sₙ) to get a monotone family spanning `X`. let T := Accumulate S.set - have T_mem : ∀ n, T n ∈ 𝒜 := fun n ↦ by + have T_mem (n) : T n ∈ 𝒜 := by simpa using h𝒜.biUnion_mem {k | k ≤ n}.toFinset (fun k _ ↦ S.set_mem k) - have T_finite : ∀ n, μ (T n) < ∞ := fun n ↦ by + have T_finite (n) : μ (T n) < ∞ := by simpa using measure_biUnion_lt_top {k | k ≤ n}.toFinset.finite_toSet (fun k _ ↦ S.finite k) have T_spanning : ⋃ n, T n = univ := S.spanning ▸ iUnion_accumulate -- We use the fact that we already know this is true for finite measures. As `⋃ n, T n = X`, -- we have that `μ ((T n) ∩ s) ⟶ μ s`. - intro s ms hμs ε ε_pos have mono : Monotone (fun n ↦ (T n) ∩ s) := fun m n hmn ↦ inter_subset_inter_left s (biUnion_subset_biUnion_left fun k hkm ↦ Nat.le_trans hkm hmn) have := tendsto_measure_iUnion_atTop (μ := μ) mono @@ -423,15 +419,14 @@ instance Lp.SecondCountableTopology [IsSeparable μ] [TopologicalSpace.Separable -- constant indicators with support in `𝒜₀`, and is denoted `D`. To make manipulations easier, -- we define the function `key` which given an integer `n` and two families of `n` elements -- in `u` and `𝒜₀` associates the corresponding sum. - let key : (n : ℕ) → (Fin n → u) → (Fin n → 𝒜₀) → (Lp E p μ) := - fun n d s ↦ ∑ i, indicatorConstLp p (h𝒜₀.measurable (s i) (Subtype.mem (s i))) - (s i).2.2 (d i : E) + let key (n : ℕ) (d : Fin n → u) (s : Fin n → 𝒜₀) : (Lp E p μ) := + ∑ i, indicatorConstLp p (h𝒜₀.measurable (s i) (Subtype.mem (s i))) (s i).2.2 (d i : E) let D := {s : Lp E p μ | ∃ n d t, s = key n d t} refine ⟨D, ?_, ?_⟩ · -- Countability directly follows from countability of `u` and `𝒜₀`. The function `f` below -- is the uncurryfied version of `key`, which is easier to manipulate as countability of the -- domain is automatically infered. - let f : (Σ n : ℕ, (Fin n → u) × (Fin n → 𝒜₀)) → Lp E p μ := fun nds ↦ key nds.1 nds.2.1 nds.2.2 + let f (nds : Σ n : ℕ, (Fin n → u) × (Fin n → 𝒜₀)) : Lp E p μ := key nds.1 nds.2.1 nds.2.2 have := count_𝒜₀.to_subtype have := countable_u.to_subtype have : D ⊆ range f := by diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 8e0ff1ca52eca..5970d297810cb 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -613,7 +613,7 @@ theorem mem_span_latticeBasis (x : (mixedSpace K)) : rfl theorem span_latticeBasis : - (Submodule.span ℤ (Set.range (latticeBasis K))) = (mixedEmbedding.integerLattice K) := + Submodule.span ℤ (Set.range (latticeBasis K)) = mixedEmbedding.integerLattice K := Submodule.ext_iff.mpr (mem_span_latticeBasis K) instance : DiscreteTopology (mixedEmbedding.integerLattice K) := by diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean index f9b153340e1c4..223002432594b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean @@ -250,22 +250,22 @@ variable (K) in /-- The set of images by `mixedEmbedding` of algebraic integers of `K` contained in the fundamental cone. -/ def integralPoint : Set (mixedSpace K) := - fundamentalCone K ∩ (mixedEmbedding.integerLattice K) + fundamentalCone K ∩ mixedEmbedding.integerLattice K theorem mem_integralPoint {a : mixedSpace K} : - a ∈ integralPoint K ↔ a ∈ fundamentalCone K ∧ ∃ x : (𝓞 K), mixedEmbedding K x = a:= by + a ∈ integralPoint K ↔ a ∈ fundamentalCone K ∧ ∃ x : 𝓞 K, mixedEmbedding K x = a := by simp only [integralPoint, Set.mem_inter_iff, SetLike.mem_coe, LinearMap.mem_range, AlgHom.toLinearMap_apply, RingHom.toIntAlgHom_coe, RingHom.coe_comp, Function.comp_apply] /-- If `a` is an integral point, then there is a *unique* algebraic integer in `𝓞 K` such that `mixedEmbedding K x = a`. -/ theorem exists_unique_preimage_of_integralPoint {a : mixedSpace K} (ha : a ∈ integralPoint K) : - ∃! x : (𝓞 K), mixedEmbedding K x = a := by + ∃! x : 𝓞 K, mixedEmbedding K x = a := by obtain ⟨_, ⟨x, rfl⟩⟩ := mem_integralPoint.mp ha refine Function.Injective.existsUnique_of_mem_range ?_ (Set.mem_range_self x) exact (mixedEmbedding_injective K).comp RingOfIntegers.coe_injective -theorem integralPoint_ne_zero (a : integralPoint K) : (a : mixedSpace K) ≠ 0 := by +theorem integralPoint_ne_zero (a : integralPoint K) : (a : mixedSpace K) ≠ 0 := by by_contra! exact a.prop.1.2 (this.symm ▸ mixedEmbedding.norm.map_zero') diff --git a/Mathlib/Probability/Distributions/Gamma.lean b/Mathlib/Probability/Distributions/Gamma.lean index 0548d04bfd3c0..562c9b3c3b6f9 100644 --- a/Mathlib/Probability/Distributions/Gamma.lean +++ b/Mathlib/Probability/Distributions/Gamma.lean @@ -30,8 +30,8 @@ open MeasureTheory Real Set Filter Topology lemma lintegral_Iic_eq_lintegral_Iio_add_Icc {y z : ℝ} (f : ℝ → ℝ≥0∞) (hzy : z ≤ y) : ∫⁻ x in Iic y, f x = (∫⁻ x in Iio z, f x) + ∫⁻ x in Icc z y, f x := by rw [← Iio_union_Icc_eq_Iic hzy, lintegral_union measurableSet_Icc] - rw [Set.disjoint_iff] - rintro x ⟨h1 : x < _, h2, _⟩ + simp_rw [Set.disjoint_iff_forall_ne, mem_Iio, mem_Icc] + intros linarith namespace ProbabilityTheory @@ -49,8 +49,9 @@ def gammaPDF (a r x : ℝ) : ℝ≥0∞ := ENNReal.ofReal (gammaPDFReal a r x) lemma gammaPDF_eq (a r x : ℝ) : - gammaPDF a r x = ENNReal.ofReal (if 0 ≤ x then - r ^ a / (Gamma a) * x ^ (a-1) * exp (-(r * x)) else 0) := rfl + gammaPDF a r x = + ENNReal.ofReal (if 0 ≤ x then r ^ a / (Gamma a) * x ^ (a-1) * exp (-(r * x)) else 0) := + rfl lemma gammaPDF_of_neg {a r x : ℝ} (hx : x < 0) : gammaPDF a r x = 0 := by simp only [gammaPDF_eq, if_neg (not_le.mpr hx), ENNReal.ofReal_zero] diff --git a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean index aaab1b9cc3f6b..f5472cc79ef0c 100644 --- a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean +++ b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean @@ -104,13 +104,10 @@ theorem le_lexOrder_iff {φ : MvPowerSeries σ R} {w : WithTop (Lex (σ →₀ intro h' have hφ : φ ≠ 0 := by rw [ne_eq, ← lexOrder_eq_top_iff_eq_zero] - intro h'' - rw [h'', ← not_le] at h' - apply h' - exact le_top + exact ne_top_of_lt h' obtain ⟨d, hd⟩ := exists_finsupp_eq_lexOrder_of_ne_zero hφ refine coeff_ne_zero_of_lexOrder hd.symm (h d ?_) - exact (lt_of_eq_of_lt hd.symm h') + rwa [← hd] theorem min_lexOrder_le {φ ψ : MvPowerSeries σ R} : min (lexOrder φ) (lexOrder ψ) ≤ lexOrder (φ + ψ) := by diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index edce0ce9c8677..f1887c8c8f025 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -323,8 +323,7 @@ theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := rw [dif_neg ane0] by_cases h : b = 0 · simp [h, lt_top_iff_ne_top] - · rw [dif_neg h] - erw [WithTop.coe_lt_coe] + · rw [dif_neg h, Nat.cast_lt] have cne0 : c ≠ 0 := by refine mt (fun con => ?_) h rw [b_eq, con, mul_zero] @@ -390,8 +389,8 @@ theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactori he ▸ e.prime_iff.1 (hp c hc), Units.map e.toMonoidHom u, by - erw [Multiset.prod_hom, ← map_mul e, h] - simp⟩ + rw [Multiset.prod_hom, toMonoidHom_eq_coe, Units.coe_map, MonoidHom.coe_coe, ← map_mul e, h, + apply_symm_apply]⟩ theorem MulEquiv.uniqueFactorizationMonoid_iff (e : α ≃* β) : UniqueFactorizationMonoid α ↔ UniqueFactorizationMonoid β := diff --git a/Mathlib/Topology/FiberPartition.lean b/Mathlib/Topology/FiberPartition.lean index c01b90f3fe77b..8e21ee229b9e1 100644 --- a/Mathlib/Topology/FiberPartition.lean +++ b/Mathlib/Topology/FiberPartition.lean @@ -29,7 +29,7 @@ variable [TopologicalSpace S] /-- The canonical map from the disjoint union induced by `f` to `S`. -/ @[simps apply] def sigmaIsoHom : C((x : Fiber f) × x.val, S) where - toFun := fun ⟨a, x⟩ ↦ x.val + toFun | ⟨a, x⟩ => x.val lemma sigmaIsoHom_inj : Function.Injective (sigmaIsoHom f) := by rintro ⟨⟨_, _, rfl⟩, ⟨_, hx⟩⟩ ⟨⟨_, _, rfl⟩, ⟨_, hy⟩⟩ h @@ -43,7 +43,7 @@ lemma sigmaIsoHom_surj : Function.Surjective (sigmaIsoHom f) := /-- The inclusion map from a component of the disjoint union induced by `f` into `S`. -/ def sigmaIncl (a : Fiber f) : C(a.val, S) where - toFun := fun x ↦ x.val + toFun x := x.val /-- The inclusion map from a fiber of a composition into the intermediate fiber. -/ def sigmaInclIncl {X : Type*} (g : Y → X) (a : Fiber (g ∘ f)) @@ -53,7 +53,7 @@ def sigmaInclIncl {X : Type*} (g : Y → X) (a : Fiber (g ∘ f)) have := x.prop simp only [sigmaIncl, ContinuousMap.coe_mk, Fiber.mem_iff_eq_image, comp_apply] at this rw [Fiber.mem_iff_eq_image, Fiber.mk_image, this, ← Fiber.map_preimage_eq_image] - rfl⟩ + simp [sigmaIncl]⟩ variable (l : LocallyConstant S Y) [CompactSpace S] From 3829434acad1d919164ed1f041a1bdb6f192ee2b Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 1 Oct 2024 08:05:20 +0000 Subject: [PATCH 073/174] chore: update Mathlib dependencies 2024-10-01 (#17322) This PR updates the Mathlib dependencies. --- lake-manifest.json | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 3e190c4de8899..c8b7ce205dedb 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "bf12ff6041cbab6eba6b54d9467baed807bb2bfd", + "rev": "4756e0fc48acce0cc808df0ad149de5973240df6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "50aaaf78b7db5bd635c19c660d59ed31b9bc9b5a", + "rev": "28fa80508edc97d96ed6342c9a771a67189e0baa", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fb7841a6f4fb389ec0e47dd4677844d49906af3c", + "rev": "e285a7ade149c551c17a4b24f127e1ef782e4bb1", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", From 809c3fb3b5c8f5d7dace56e200b426187516535a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 1 Oct 2024 08:15:40 +0000 Subject: [PATCH 074/174] chore: move to v4.12.0 (#17320) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 98556ba065e2a..89985206aca4e 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0-rc1 +leanprover/lean4:v4.12.0 From d2126e2b3ffb31758673525c892aa27e1bbd8e9a Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 1 Oct 2024 10:10:25 +0000 Subject: [PATCH 075/174] chore: update Mathlib dependencies 2024-10-01 (#17324) 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 c8b7ce205dedb..6c6e4f48601c1 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2ba60fa2c384a94735454db11a2d523612eaabff", + "rev": "4261c7c9290da3471548a327a4fa9387e9fdd684", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From 581d005cdb91f75371c63b4dd67283dc3d38f690 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 1 Oct 2024 10:43:13 +0000 Subject: [PATCH 076/174] =?UTF-8?q?feat:=20inverse=20of=20a=201=C3=971=20m?= =?UTF-8?q?atrix=20(#17298)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The statement is a little strange, so it comes with a docstring to justify it. From https://github.com/eric-wieser/lean-matrix-cookbook Co-authored-by: Eric Wieser --- Mathlib/Data/Matrix/Basic.lean | 5 +++++ .../Matrix/NonsingularInverse.lean | 22 +++++++++++++++++++ 2 files changed, 27 insertions(+) diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index dfb96ff464262..6bcfb3a9df82c 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -492,6 +492,11 @@ theorem diagonal_conjTranspose [AddMonoid α] [StarAddMonoid α] (v : n → α) rw [conjTranspose, diagonal_transpose, diagonal_map (star_zero _)] rfl +theorem diagonal_unique [Unique m] [DecidableEq m] [Zero α] (d : m → α) : + diagonal d = of fun _ _ => d default := by + ext i j + rw [Subsingleton.elim i default, Subsingleton.elim j default, diagonal_apply_eq _ _, of_apply] + section One variable [Zero α] [One α] diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 1482f940cd82d..6ef3b1d0b480f 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -470,6 +470,10 @@ theorem nonsing_inv_nonsing_inv (h : IsUnit A.det) : A⁻¹⁻¹ = A := theorem isUnit_nonsing_inv_det_iff {A : Matrix n n α} : IsUnit A⁻¹.det ↔ IsUnit A.det := by rw [Matrix.det_nonsing_inv, isUnit_ring_inverse] +@[simp] +theorem isUnit_nonsing_inv_iff {A : Matrix n n α} : IsUnit A⁻¹ ↔ IsUnit A := by + simp_rw [isUnit_iff_isUnit_det, isUnit_nonsing_inv_det_iff] + -- `IsUnit.invertible` lifts the proposition `IsUnit A` to a constructive inverse of `A`. /-- A version of `Matrix.invertibleOfDetInvertible` with the inverse defeq to `A⁻¹` that is therefore noncomputable. -/ @@ -607,6 +611,24 @@ theorem inv_diagonal (v : n → α) : (diagonal v)⁻¹ = diagonal (Ring.inverse end Diagonal +/-- The inverse of a 1×1 or 0×0 matrix is always diagonal. + +While we could write this as `of fun _ _ => Ring.inverse (A default default)` on the RHS, this is +less useful because: + +* It wouldn't work for 0×0 matrices. +* More things are true about diagonal matrices than constant matrices, and so more lemmas exist. + +`Matrix.diagonal_unique` can be used to reach this form, while `Ring.inverse_eq_inv` can be used +to replace `Ring.inverse` with `⁻¹`. +-/ +@[simp] +theorem inv_subsingleton [Subsingleton m] [Fintype m] [DecidableEq m] (A : Matrix m m α) : + A⁻¹ = diagonal fun i => Ring.inverse (A i i) := by + rw [inv_def, adjugate_subsingleton, smul_one_eq_diagonal] + congr! with i + exact det_eq_elem_of_subsingleton _ _ + section Woodbury variable [Fintype m] [DecidableEq m] From c24c460adbd11a4ffd84cf75872e20bcf167bdbe Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 1 Oct 2024 11:19:30 +0000 Subject: [PATCH 077/174] =?UTF-8?q?feat:=20`Irrational=20=CF=80`=20(#17103?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This contains a port of @b-mehta's https://github.com/leanprover-community/mathlib3/commits/irrational-pi branch, with the conversion to Lean 4 performed by @b-mehta with his google hat on. My contributions here are minimal! Co-authored-by: Bhavik Mehta Co-authored-by: Bhavik Mehta Co-authored-by: Bhavik Mehta Co-authored-by: Eric Wieser Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 + .../Polynomial/Degree/Definitions.lean | 6 + Mathlib/Data/Real/Pi/Irrational.lean | 305 ++++++++++++++++++ .../Integral/IntervalIntegral.lean | 6 + 4 files changed, 318 insertions(+) create mode 100644 Mathlib/Data/Real/Pi/Irrational.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8211ab834a536..1ed36f8f87b9c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2518,6 +2518,7 @@ import Mathlib.Data.Real.GoldenRatio import Mathlib.Data.Real.Hyperreal import Mathlib.Data.Real.Irrational import Mathlib.Data.Real.Pi.Bounds +import Mathlib.Data.Real.Pi.Irrational import Mathlib.Data.Real.Pi.Leibniz import Mathlib.Data.Real.Pi.Wallis import Mathlib.Data.Real.Pointwise diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 1d9ab3832d4fb..3a3fe3f22832f 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -237,6 +237,12 @@ theorem natDegree_natCast (n : ℕ) : natDegree (n : R[X]) = 0 := by @[deprecated (since := "2024-04-17")] alias natDegree_nat_cast := natDegree_natCast +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem natDegree_ofNat (n : ℕ) [Nat.AtLeastTwo n] : + natDegree (no_index (OfNat.ofNat n : R[X])) = 0 := + natDegree_natCast _ + theorem degree_natCast_le (n : ℕ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp) @[deprecated (since := "2024-04-17")] diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean new file mode 100644 index 0000000000000..8736c6f964f66 --- /dev/null +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -0,0 +1,305 @@ +/- +Copyright (c) 2022 Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bhavik Mehta +-/ +import Mathlib.Analysis.SpecialFunctions.Integrals +import Mathlib.Data.Real.Irrational + +/-! +# `Real.pi` is irrational + +The main result of this file is `irrational_pi`. + +The proof is adapted from https://en.wikipedia.org/wiki/Proof_that_%CF%80_is_irrational#Cartwright's_proof. + +The proof idea is as follows. +* Define a sequence of integrals `I n θ = ∫ x in (-1)..1, (1 - x ^ 2) ^ n * cos (x * θ)`. +* Give a recursion formula for `I (n + 2) θ * θ ^ 2` in terms of `I n θ` and `I (n + 1) θ`. + Note we do not find it helpful to define `J` as in the above proof, and instead work directly + with `I`. +* Define polynomials with integer coefficients `sinPoly n` and `cosPoly n` such that + `I n θ * θ ^ (2 * n + 1) = n ! * (sinPoly n θ * sin θ + cosPoly n θ * cos θ)`. + Note that in the informal proof, these polynomials are not defined explicitly, but we find it + useful to define them by recursion. +* Show that both these polynomials have degree bounded by `n`. +* Show that `0 < I n (π / 2) ≤ 2` for all `n`. +* Now we can finish: if `π / 2` is rational, write it as `a / b` with `a, b > 0`. Then + `b ^ (2 * n + 1) * sinPoly n (a / b)` is a positive integer by the degree bound. But it is equal + to `a ^ (2 * n + 1) / n ! * I n (π / 2) ≤ 2 * a * (2 * n + 1) / n !`, which converges to 0 as + `n → ∞`. + +-/ + +noncomputable section + +open intervalIntegral MeasureTheory.MeasureSpace Set Polynomial Real +open scoped Nat + +/-- The sequence of integrals used for Cartwright's proof of irrationality of `π`. -/ +private def I (n : ℕ) (θ : ℝ) : ℝ := ∫ x in (-1)..1, (1 - x ^ 2) ^ n * cos (x * θ) + +variable {n : ℕ} {θ : ℝ} + +private lemma I_zero : I 0 θ * θ = 2 * sin θ := by + rw [mul_comm, I] + simp [mul_integral_comp_mul_right, two_mul] + +/-- +Auxiliary for the proof that `π` is irrational. +While it is most natural to give the recursive formula for `I (n + 2) θ`, as well as give the second +base case of `I 1 θ`, it is in fact more convenient to give the recursive formula for `I (n + 1) θ` +in terms of `I n θ` and `I (n - 1) θ` (note the natural subtraction!). +Despite the usually inconvenient subtraction, this in fact allows deducing both of the above facts +with significantly fewer analysis computations. +In addition, note the `0 ^ n` on the right hand side - this is intentional, and again allows +combining the proof of the "usual" recursion formula and the base case `I 1 θ`. +-/ +private lemma recursion' (n : ℕ) : + I (n + 1) θ * θ ^ 2 = - (2 * 2 * ((n + 1) * (0 ^ n * cos θ))) + + 2 * (n + 1) * (2 * n + 1) * I n θ - 4 * (n + 1) * n * I (n - 1) θ := by + rw [I] + let f (x : ℝ) : ℝ := 1 - x ^ 2 + let u₁ (x : ℝ) : ℝ := f x ^ (n + 1) + let u₁' (x : ℝ) : ℝ := - (2 * (n + 1) * x * f x ^ n) + let v₁ (x : ℝ) : ℝ := sin (x * θ) + let v₁' (x : ℝ) : ℝ := cos (x * θ) * θ + let u₂ (x : ℝ) : ℝ := x * (f x) ^ n + let u₂' (x : ℝ) : ℝ := (f x) ^ n - 2 * n * x ^ 2 * (f x) ^ (n - 1) + let v₂ (x : ℝ) : ℝ := cos (x * θ) + let v₂' (x : ℝ) : ℝ := -sin (x * θ) * θ + have hfd : Continuous f := by fun_prop + have hu₁d : Continuous u₁' := by fun_prop + have hv₁d : Continuous v₁' := by fun_prop + have hu₂d : Continuous u₂' := by fun_prop + have hv₂d : Continuous v₂' := by fun_prop + have hu₁_eval_one : u₁ 1 = 0 := by simp only [u₁, f]; simp + have hu₁_eval_neg_one : u₁ (-1) = 0 := by simp only [u₁, f]; simp + have t : u₂ 1 * v₂ 1 - u₂ (-1) * v₂ (-1) = 2 * (0 ^ n * cos θ) := by simp [u₂, v₂, f, ← two_mul] + have hf (x) : HasDerivAt f (- 2 * x) x := by + convert (hasDerivAt_pow 2 x).const_sub 1 using 1 + simp + have hu₁ (x) : HasDerivAt u₁ (u₁' x) x := by + convert (hf x).pow _ using 1 + simp only [Nat.add_succ_sub_one, u₁', Nat.cast_add_one] + ring + have hv₁ (x) : HasDerivAt v₁ (v₁' x) x := (hasDerivAt_mul_const θ).sin + have hu₂ (x) : HasDerivAt u₂ (u₂' x) x := by + convert (hasDerivAt_id' x).mul ((hf x).pow _) using 1 + simp only [u₂'] + ring + have hv₂ (x) : HasDerivAt v₂ (v₂' x) x := (hasDerivAt_mul_const θ).cos + convert_to (∫ (x : ℝ) in (-1)..1, u₁ x * v₁' x) * θ = _ using 1 + · simp_rw [u₁, v₁', ← intervalIntegral.integral_mul_const, sq θ, mul_assoc] + rw [integral_mul_deriv_eq_deriv_mul (fun x _ => hu₁ x) (fun x _ => hv₁ x) + (hu₁d.intervalIntegrable _ _) (hv₁d.intervalIntegrable _ _), hu₁_eval_one, hu₁_eval_neg_one, + zero_mul, zero_mul, sub_zero, zero_sub, ← integral_neg, ← integral_mul_const] + convert_to ((-2 : ℝ) * (n + 1)) * ∫ (x : ℝ) in (-1)..1, (u₂ x * v₂' x) = _ using 1 + · rw [← integral_const_mul] + congr 1 with x + dsimp [u₁', v₁, u₂, v₂'] + ring + rw [integral_mul_deriv_eq_deriv_mul (fun x _ => hu₂ x) (fun x _ => hv₂ x) + (hu₂d.intervalIntegrable _ _) (hv₂d.intervalIntegrable _ _), + mul_sub, t, neg_mul, neg_mul, neg_mul, sub_neg_eq_add] + have (x) : u₂' x = (2 * n + 1) * f x ^ n - 2 * n * f x ^ (n - 1) := by + cases n with + | zero => simp [u₂'] + | succ n => ring! + simp_rw [this, sub_mul, mul_assoc _ _ (v₂ _)] + have : Continuous v₂ := by fun_prop + rw [mul_mul_mul_comm, integral_sub, mul_sub, add_sub_assoc] + · congr 1 + simp_rw [integral_const_mul] + ring! + all_goals exact Continuous.intervalIntegrable (by fun_prop) _ _ + +/-- +Auxiliary for the proof that `π` is irrational. +The recursive formula for `I (n + 2) θ * θ ^ 2` in terms of `I n θ` and `I (n + 1) θ`. +-/ +private lemma recursion (n : ℕ) : + I (n + 2) θ * θ ^ 2 = + 2 * (n + 2) * (2 * n + 3) * I (n + 1) θ - 4 * (n + 2) * (n + 1) * I n θ := by + rw [recursion' (n + 1)] + simp + ring! + +/-- +Auxiliary for the proof that `π` is irrational. +The second base case for the induction on `n`, giving an explicit formula for `I 1 θ`. +-/ +private lemma I_one : I 1 θ * θ ^ 3 = 4 * sin θ - 4 * θ * cos θ := by + rw [_root_.pow_succ, ← mul_assoc, recursion' 0, sub_mul, add_mul, mul_assoc _ (I 0 θ), I_zero] + ring + +/-- +Auxiliary for the proof that `π` is irrational. +The first of the two integer-coefficient polynomials that describe the behaviour of the +sequence of integrals `I`. +While not given in the informal proof, these are easy to deduce from the recursion formulae. +-/ +private def sinPoly : ℕ → ℤ[X] + | 0 => C 2 + | 1 => C 4 + | (n+2) => ((2 : ℤ) * (2 * n + 3)) • sinPoly (n + 1) + monomial 2 (-4) * sinPoly n + +/-- +Auxiliary for the proof that `π` is irrational. +The second of the two integer-coefficient polynomials that describe the behaviour of the +sequence of integrals `I`. +While not given in the informal proof, these are easy to deduce from the recursion formulae. +-/ +private def cosPoly : ℕ → ℤ[X] + | 0 => 0 + | 1 => monomial 1 (-4) + | (n+2) => ((2 : ℤ) * (2 * n + 3)) • cosPoly (n + 1) + monomial 2 (-4) * cosPoly n + +/-- +Auxiliary for the proof that `π` is irrational. +Prove a degree bound for `sinPoly n` by induction. Note this is where we find the value in an +explicit description of `sinPoly`. +-/ +private lemma sinPoly_natDegree_le : ∀ n : ℕ, (sinPoly n).natDegree ≤ n + | 0 => by simp [sinPoly] + | 1 => by simp only [natDegree_C, mul_one, zero_le', sinPoly] + | n + 2 => by + rw [sinPoly] + refine natDegree_add_le_of_degree_le ((natDegree_smul_le _ _).trans ?_) ?_ + · exact (sinPoly_natDegree_le (n + 1)).trans (by simp) + refine natDegree_mul_le.trans ?_ + simpa [add_comm 2] using sinPoly_natDegree_le n + +/-- +Auxiliary for the proof that `π` is irrational. +Prove a degree bound for `cosPoly n` by induction. Note this is where we find the value in an +explicit description of `cosPoly`. +-/ +private lemma cosPoly_natDegree_le : ∀ n : ℕ, (cosPoly n).natDegree ≤ n + | 0 => by simp [cosPoly] + | 1 => (natDegree_monomial_le _).trans (by simp) + | n + 2 => by + rw [cosPoly] + refine natDegree_add_le_of_degree_le ((natDegree_smul_le _ _).trans ?_) ?_ + · exact (cosPoly_natDegree_le (n + 1)).trans (by simp) + exact natDegree_mul_le.trans (by simp [add_comm 2, cosPoly_natDegree_le n]) + +/-- +Auxiliary for the proof that `π` is irrational. +The key lemma: the sequence of integrals `I` can be written as a linear combination of `sin` and +`cos`, with coefficients given by the polynomials `sinPoly` and `cosPoly`. +-/ +private lemma sinPoly_add_cosPoly_eval (θ : ℝ) : + ∀ n : ℕ, + I n θ * θ ^ (2 * n + 1) = n ! * ((sinPoly n).eval₂ (Int.castRingHom _) θ * sin θ + + (cosPoly n).eval₂ (Int.castRingHom _) θ * cos θ) + | 0 => by simp [sinPoly, cosPoly, I_zero] + | 1 => by simp [I_one, sinPoly, cosPoly, sub_eq_add_neg] + | n + 2 => by + calc I (n + 2) θ * θ ^ (2 * (n + 2) + 1) = I (n + 2) θ * θ ^ 2 * θ ^ (2 * n + 3) := by ring + _ = 2 * (n + 2) * (2 * n + 3) * (I (n + 1) θ * θ ^ (2 * (n + 1) + 1)) - + 4 * (n + 2) * (n + 1) * θ ^ 2 * (I n θ * θ ^ (2 * n + 1)) := by rw [recursion]; ring + _ = _ := by simp [sinPoly_add_cosPoly_eval, sinPoly, cosPoly, Nat.factorial_succ]; ring + +/-- +Auxiliary for the proof that `π` is irrational. +For a polynomial `p` with natural degree `≤ k` and integer coefficients, evaluating `p` at a +rational `a / b` gives a rational of the form `z / b ^ k`. +TODO: should this be moved elsewhere? It uses none of the pi-specific definitions. +-/ +private lemma is_integer {p : ℤ[X]} (a b : ℤ) {k : ℕ} (hp : p.natDegree ≤ k) : + ∃ z : ℤ, p.eval₂ (Int.castRingHom ℝ) (a / b) * b ^ k = z := by + rcases eq_or_ne b 0 with rfl | hb + · rcases k.eq_zero_or_pos with rfl | hk + · exact ⟨p.coeff 0, by simp⟩ + exact ⟨0, by simp [hk.ne']⟩ + refine ⟨∑ i in p.support, p.coeff i * a ^ i * b ^ (k - i), ?_⟩ + conv => lhs; rw [← sum_monomial_eq p] + rw [eval₂_sum, sum, Finset.sum_mul, Int.cast_sum] + simp only [eval₂_monomial, eq_intCast, div_pow, Int.cast_mul, Int.cast_pow] + refine Finset.sum_congr rfl (fun i hi => ?_) + have ik := (le_natDegree_of_mem_supp i hi).trans hp + rw [mul_assoc, div_mul_comm, ← Int.cast_pow, ← Int.cast_pow, ← Int.cast_pow, + ← pow_sub_mul_pow b ik, ← Int.cast_div_charZero, Int.mul_ediv_cancel _ (pow_ne_zero _ hb), + ← mul_assoc, mul_right_comm, ← Int.cast_pow] + exact dvd_mul_left _ _ + +open Filter + +/-- +Auxiliary for the proof that `π` is irrational. +The integrand in the definition of `I` is nonnegative and takes a positive value at least one point, +so the integral is positive. +-/ +private lemma I_pos : 0 < I n (π / 2) := by + refine integral_pos (by norm_num) (Continuous.continuousOn (by continuity)) ?_ ⟨0, by simp⟩ + refine fun x hx => mul_nonneg (pow_nonneg ?_ _) ?_ + · rw [sub_nonneg, sq_le_one_iff_abs_le_one, abs_le] + exact ⟨hx.1.le, hx.2⟩ + refine cos_nonneg_of_neg_pi_div_two_le_of_le ?_ ?_ <;> + nlinarith [hx.1, hx.2, pi_pos] + +/-- +Auxiliary for the proof that `π` is irrational. +The integrand in the definition of `I` is bounded by 1 and the interval has length 2, so the +integral is bounded above by `2`. +-/ +private lemma I_le (n : ℕ) : I n (π / 2) ≤ 2 := by + rw [← norm_of_nonneg I_pos.le] + refine (norm_integral_le_of_norm_le_const ?_).trans (show (1 : ℝ) * _ ≤ _ by norm_num) + intros x hx + simp only [uIoc_of_le, neg_le_self_iff, zero_le_one, mem_Ioc] at hx + rw [norm_eq_abs, abs_mul, abs_pow] + refine mul_le_one (pow_le_one _ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) + rw [abs_le] + constructor <;> nlinarith + +/-- +Auxiliary for the proof that `π` is irrational. +For any real `a`, we have that `a ^ (2n+1) / n!` tends to `0` as `n → ∞`. This is just a +reformulation of tendsto_pow_div_factorial_atTop, which asserts the same for `a ^ n / n!` +-/ +private lemma tendsto_pow_div_factorial_at_top_aux (a : ℝ) : + Tendsto (fun n => (a : ℝ) ^ (2 * n + 1) / n !) atTop (nhds 0) := by + rw [← mul_zero a] + refine ((tendsto_pow_div_factorial_atTop (a ^ 2)).const_mul a).congr (fun x => ?_) + rw [← pow_mul, mul_div_assoc', _root_.pow_succ'] + +/-- If `x` is rational, it can be written as `a / b` with `a : ℤ` and `b : ℕ` satisfying `b > 0`. -/ +private lemma not_irrational_exists_rep {x : ℝ} : + ¬Irrational x → ∃ (a : ℤ) (b : ℕ), 0 < b ∧ x = a / b := by + rw [Irrational, not_not, mem_range] + rintro ⟨q, rfl⟩ + exact ⟨q.num, q.den, q.pos, by exact_mod_cast (Rat.num_div_den _).symm⟩ + +@[simp] theorem irrational_pi : Irrational π := by + apply Irrational.of_div_nat 2 + rw [Nat.cast_two] + by_contra h' + obtain ⟨a, b, hb, h⟩ := not_irrational_exists_rep h' + have ha : (0 : ℝ) < a := by + have : 0 < (a : ℝ) / b := h ▸ pi_div_two_pos + rwa [lt_div_iff (by positivity), zero_mul] at this + have k (n : ℕ) : 0 < (a : ℝ) ^ (2 * n + 1) / n ! := by positivity + have j : ∀ᶠ n : ℕ in atTop, (a : ℝ) ^ (2 * n + 1) / n ! * I n (π / 2) < 1 := by + have := eventually_lt_of_tendsto_lt (show (0 : ℝ) < 1 / 2 by norm_num) + (tendsto_pow_div_factorial_at_top_aux a) + filter_upwards [this] with n hn + rw [lt_div_iff (zero_lt_two : (0 : ℝ) < 2)] at hn + exact hn.trans_le' (mul_le_mul_of_nonneg_left (I_le _) (by positivity)) + obtain ⟨n, hn⟩ := j.exists + have hn' : 0 < a ^ (2 * n + 1) / n ! * I n (π / 2) := mul_pos (k _) I_pos + obtain ⟨z, hz⟩ : ∃ z : ℤ, (sinPoly n).eval₂ (Int.castRingHom ℝ) (a / b) * b ^ (2 * n + 1) = z := + is_integer a b ((sinPoly_natDegree_le _).trans (by linarith)) + have e := sinPoly_add_cosPoly_eval (π / 2) n + rw [cos_pi_div_two, sin_pi_div_two, mul_zero, mul_one, add_zero] at e + have : a ^ (2 * n + 1) / n ! * I n (π / 2) = + eval₂ (Int.castRingHom ℝ) (π / 2) (sinPoly n) * b ^ (2 * n + 1) := by + nth_rw 2 [h] at e + field_simp at e ⊢ + linear_combination e + have : (0 : ℝ) < z ∧ (z : ℝ) < 1 := by simp [← hz, ← h, ← this, hn', hn] + norm_cast at this + omega + +end diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index fc342e5424d56..d7473d6e2d8df 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -1021,6 +1021,12 @@ theorem abs_integral_le_integral_abs (hab : a ≤ b) : |∫ x in a..b, f x ∂μ| ≤ ∫ x in a..b, |f x| ∂μ := by simpa only [← Real.norm_eq_abs] using norm_integral_le_integral_norm hab +lemma integral_pos (hab : a < b) + (hfc : ContinuousOn f (Icc a b)) (hle : ∀ x ∈ Ioc a b, 0 ≤ f x) (hlt : ∃ c ∈ Icc a b, 0 < f c) : + 0 < ∫ x in a..b, f x := + (integral_lt_integral_of_continuousOn_of_le_of_exists_lt hab + continuousOn_const hfc hle hlt).trans_eq' (by simp) + section Mono theorem integral_mono_interval {c d} (hca : c ≤ a) (hab : a ≤ b) (hbd : b ≤ d) From 4ae56c115bda14b18fe7ac68347a73a4bb7f6bac Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 1 Oct 2024 11:19:31 +0000 Subject: [PATCH 078/174] chore: rename `unique_diff` to `uniqueDiffOn` in theorems (#17291) --- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 6 +-- Mathlib/Geometry/Manifold/Diffeomorph.lean | 2 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 4 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 2 +- .../Manifold/MFDeriv/SpecificFunctions.lean | 2 +- .../Manifold/SmoothManifoldWithCorners.lean | 38 ++++++++++++------- .../Manifold/VectorBundle/Tangent.lean | 6 +-- 7 files changed, 36 insertions(+), 24 deletions(-) diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 3dba3c04792f4..0b4140f98c392 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -96,7 +96,7 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → rw [contMDiffAt_iff] at hf hg simp_rw [Function.comp_def, uncurry, extChartAt_prod, PartialEquiv.prod_coe_symm, ModelWithCorners.range_prod] at hf ⊢ - refine ContDiffWithinAt.fderivWithin ?_ hg.2 I.unique_diff hmn (mem_range_self _) ?_ + refine ContDiffWithinAt.fderivWithin ?_ hg.2 I.uniqueDiffOn hmn (mem_range_self _) ?_ · simp_rw [extChartAt_to_inv]; exact hf.2 · rw [← image_subset_iff] rintro _ ⟨x, -, rfl⟩ @@ -165,7 +165,7 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → PartialEquiv.mem_symm_trans_source _ (mem_extChartAt_source I' (f x₂ (g x₂))) h3x₂).differentiableWithinAt le_top have h3f := (h2x₂.mdifferentiableAt le_rfl).differentiableWithinAt_writtenInExtChartAt - refine fderivWithin.comp₃ _ hI' h3f hI ?_ ?_ ?_ ?_ (I.unique_diff _ <| mem_range_self _) + 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, @@ -577,7 +577,7 @@ theorem tangentMap_tangentBundle_pure [Is : SmoothManifoldWithCorners I M] (p : · simp · exact differentiableAt_id' · exact differentiableAt_const _ - · exact ModelWithCorners.unique_diff_at_image I + · exact ModelWithCorners.uniqueDiffWithinAt_image I · exact differentiableAt_id'.prod (differentiableAt_const _) simp (config := { unfoldPartialApp := true }) only [Bundle.zeroSection, tangentMap, mfderiv, A, if_pos, chartAt, FiberBundle.chartedSpace_chartAt, TangentBundle.trivializationAt_apply, diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 2c6f481393275..07f8ab47bb49f 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -447,7 +447,7 @@ variable (I) (e : E ≃ₘ[𝕜] E') def transDiffeomorph (I : ModelWithCorners 𝕜 E H) (e : E ≃ₘ[𝕜] E') : ModelWithCorners 𝕜 E' H where toPartialEquiv := I.toPartialEquiv.trans e.toEquiv.toPartialEquiv source_eq := by simp - unique_diff' := by simp [range_comp e, I.unique_diff] + uniqueDiffOn' := by simp [range_comp e, I.uniqueDiffOn] continuous_toFun := e.continuous.comp I.continuous continuous_invFun := I.continuous_symm.comp e.symm.continuous diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 5a704749a28a5..43cbf79926a58 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -136,7 +136,7 @@ def modelWithCornersEuclideanHalfSpace (n : ℕ) [NeZero n] : exact ⟨max_eq_left xprop, fun i _ => rfl⟩ right_inv' x hx := update_eq_iff.2 ⟨max_eq_left hx, fun i _ => rfl⟩ source_eq := rfl - unique_diff' := by + uniqueDiffOn' := by have : UniqueDiffOn ℝ _ := UniqueDiffOn.pi (Fin n) (fun _ => ℝ) _ _ fun i (_ : i ∈ ({0} : Set (Fin n))) => uniqueDiffOn_Ici 0 @@ -159,7 +159,7 @@ def modelWithCornersEuclideanQuadrant (n : ℕ) : left_inv' x _ := by ext i; simp only [Subtype.coe_mk, x.2 i, max_eq_left] right_inv' x hx := by ext1 i; simp only [hx i, max_eq_left] source_eq := rfl - unique_diff' := by + uniqueDiffOn' := by have this : UniqueDiffOn ℝ _ := UniqueDiffOn.univ_pi (Fin n) (fun _ => ℝ) _ fun _ => uniqueDiffOn_Ici 0 simpa only [pi_univ_Ici] using this diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index f9b96b110f3fc..a5e16af8e340f 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -44,7 +44,7 @@ variable theorem uniqueMDiffWithinAt_univ : UniqueMDiffWithinAt I univ x := by unfold UniqueMDiffWithinAt simp only [preimage_univ, univ_inter] - exact I.unique_diff _ (mem_range_self _) + exact I.uniqueDiffOn _ (mem_range_self _) variable {I} diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index 75f5779ee5a46..5e0909b901206 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -324,7 +324,7 @@ theorem MDifferentiableAt.mfderiv_prod {f : M → M'} {g : M → M''} {x : M} classical simp_rw [mfderiv, if_pos (hf.prod_mk hg), if_pos hf, if_pos hg] exact hf.differentiableWithinAt_writtenInExtChartAt.fderivWithin_prod - hg.differentiableWithinAt_writtenInExtChartAt (I.unique_diff _ (mem_range_self _)) + hg.differentiableWithinAt_writtenInExtChartAt (I.uniqueDiffOn _ (mem_range_self _)) variable (I I' I'') diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 5ed443618c56c..9f6ab2d09bb29 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -138,7 +138,7 @@ structure ModelWithCorners (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Ty [NormedAddCommGroup E] [NormedSpace 𝕜 E] (H : Type*) [TopologicalSpace H] extends PartialEquiv H E where source_eq : source = univ - unique_diff' : UniqueDiffOn 𝕜 toPartialEquiv.target + uniqueDiffOn' : UniqueDiffOn 𝕜 toPartialEquiv.target continuous_toFun : Continuous toFun := by continuity continuous_invFun : Continuous invFun := by continuity @@ -149,7 +149,7 @@ def modelWithCornersSelf (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type [NormedAddCommGroup E] [NormedSpace 𝕜 E] : ModelWithCorners 𝕜 E E where toPartialEquiv := PartialEquiv.refl E source_eq := rfl - unique_diff' := uniqueDiffOn_univ + uniqueDiffOn' := uniqueDiffOn_univ continuous_toFun := continuous_id continuous_invFun := continuous_id @@ -236,8 +236,11 @@ theorem target_eq : I.target = range (I : H → E) := by rw [← image_univ, ← I.source_eq] exact I.image_source_eq_target.symm -protected theorem unique_diff : UniqueDiffOn 𝕜 (range I) := - I.target_eq ▸ I.unique_diff' +protected theorem uniqueDiffOn : UniqueDiffOn 𝕜 (range I) := + I.target_eq ▸ I.uniqueDiffOn' + +@[deprecated (since := "2024-09-30")] +protected alias unique_diff := ModelWithCorners.uniqueDiffOn @[simp, mfld_simps] protected theorem left_inv (x : H) : I.symm (I x) = x := by refine I.left_inv' ?_; simp @@ -290,17 +293,26 @@ theorem symm_map_nhdsWithin_image {x : H} {s : Set H} : map I.symm (𝓝[I '' s] theorem symm_map_nhdsWithin_range (x : H) : map I.symm (𝓝[range I] I x) = 𝓝 x := by rw [← I.map_nhds_eq, map_map, I.symm_comp_self, map_id] -theorem unique_diff_preimage {s : Set H} (hs : IsOpen s) : +theorem uniqueDiffOn_preimage {s : Set H} (hs : IsOpen s) : UniqueDiffOn 𝕜 (I.symm ⁻¹' s ∩ range I) := by rw [inter_comm] - exact I.unique_diff.inter (hs.preimage I.continuous_invFun) + exact I.uniqueDiffOn.inter (hs.preimage I.continuous_invFun) + +@[deprecated (since := "2024-09-30")] +alias unique_diff_preimage := uniqueDiffOn_preimage -theorem unique_diff_preimage_source {β : Type*} [TopologicalSpace β] {e : PartialHomeomorph H β} : +theorem uniqueDiffOn_preimage_source {β : Type*} [TopologicalSpace β] {e : PartialHomeomorph H β} : UniqueDiffOn 𝕜 (I.symm ⁻¹' e.source ∩ range I) := - I.unique_diff_preimage e.open_source + I.uniqueDiffOn_preimage e.open_source + +@[deprecated (since := "2024-09-30")] +alias unique_diff_preimage_source := uniqueDiffOn_preimage_source + +theorem uniqueDiffWithinAt_image {x : H} : UniqueDiffWithinAt 𝕜 (range I) (I x) := + I.uniqueDiffOn _ (mem_range_self _) -theorem unique_diff_at_image {x : H} : UniqueDiffWithinAt 𝕜 (range I) (I x) := - I.unique_diff _ (mem_range_self _) +@[deprecated (since := "2024-09-30")] +alias unique_diff_at_image := uniqueDiffWithinAt_image theorem symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] {f : H → X} {s : Set H} {x : H} : @@ -369,7 +381,7 @@ def ModelWithCorners.prod {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Ty invFun := fun x => (I.symm x.1, I'.symm x.2) source := { x | x.1 ∈ I.source ∧ x.2 ∈ I'.source } source_eq := by simp only [setOf_true, mfld_simps] - unique_diff' := I.unique_diff'.prod I'.unique_diff' + uniqueDiffOn' := I.uniqueDiffOn'.prod I'.uniqueDiffOn' continuous_toFun := I.continuous_toFun.prod_map I'.continuous_toFun continuous_invFun := I.continuous_invFun.prod_map I'.continuous_invFun } @@ -382,7 +394,7 @@ def ModelWithCorners.pi {𝕜 : Type u} [NontriviallyNormedField 𝕜] {ι : Typ ModelWithCorners 𝕜 (∀ i, E i) (ModelPi H) where toPartialEquiv := PartialEquiv.pi fun i => (I i).toPartialEquiv source_eq := by simp only [pi_univ, mfld_simps] - unique_diff' := UniqueDiffOn.pi ι E _ _ fun i _ => (I i).unique_diff' + uniqueDiffOn' := UniqueDiffOn.pi ι E _ _ fun i _ => (I i).uniqueDiffOn' continuous_toFun := continuous_pi fun i => (I i).continuous.comp (continuous_apply i) continuous_invFun := continuous_pi fun i => (I i).continuous_symm.comp (continuous_apply i) @@ -1057,7 +1069,7 @@ theorem extChartAt_target (x : M) : theorem uniqueDiffOn_extChartAt_target (x : M) : UniqueDiffOn 𝕜 (extChartAt I x).target := by rw [extChartAt_target] - exact I.unique_diff_preimage (chartAt H x).open_target + exact I.uniqueDiffOn_preimage (chartAt H x).open_target theorem uniqueDiffWithinAt_extChartAt_target (x : M) : UniqueDiffWithinAt 𝕜 (extChartAt I x).target (extChartAt I x x) := diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index fed71bc4299df..df7663e67c7ac 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -55,7 +55,7 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : have h : ((i.1.extend I).symm ≫ j.1.extend I).source ⊆ range I := by rw [i.1.extend_coord_change_source]; apply image_subset_range intro x hx - refine (ContDiffWithinAt.fderivWithin_right ?_ I.unique_diff le_top <| h hx).mono h + 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 @@ -84,7 +84,7 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where coordChange_self i x hx v := by simp only rw [Filter.EventuallyEq.fderivWithin_eq, fderivWithin_id', ContinuousLinearMap.id_apply] - · exact I.unique_diff_at_image + · exact I.uniqueDiffWithinAt_image · filter_upwards [i.1.extend_target_mem_nhdsWithin I hx] with y hy exact (i.1.extend I).right_inv hy · simp_rw [Function.comp_apply, i.1.extend_left_inv I hx] @@ -105,7 +105,7 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where · exact (contDiffWithinAt_extend_coord_change' I (subset_maximalAtlas I j.2) (subset_maximalAtlas I i.2) hxj hxi).differentiableWithinAt le_top · intro x _; exact mem_range_self _ - · exact I.unique_diff_at_image + · exact I.uniqueDiffWithinAt_image · rw [Function.comp_apply, i.1.extend_left_inv I hxi] -- Porting note: moved to a separate `simp high` lemma b/c `simp` can simplify the LHS From 69c3e9253dfe613022a408aaa4af323855209d6e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 11:49:13 +0000 Subject: [PATCH 079/174] feat: Convex functions are continuous (#14999) Prove that convex functions from an open set in a finite dimensional real normed space are locally Lipschitz on that set. --- Mathlib.lean | 1 + Mathlib/Analysis/Convex/Continuous.lean | 232 ++++++++++++++++++++++++ Mathlib/Analysis/Convex/Intrinsic.lean | 27 --- Mathlib/Topology/ContinuousOn.lean | 2 +- docs/undergrad.yaml | 4 +- 5 files changed, 237 insertions(+), 29 deletions(-) create mode 100644 Mathlib/Analysis/Convex/Continuous.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1ed36f8f87b9c..b7d0a3ed8d92f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1097,6 +1097,7 @@ import Mathlib.Analysis.Convex.Cone.Extension import Mathlib.Analysis.Convex.Cone.InnerDual import Mathlib.Analysis.Convex.Cone.Pointed import Mathlib.Analysis.Convex.Cone.Proper +import Mathlib.Analysis.Convex.Continuous import Mathlib.Analysis.Convex.Contractible import Mathlib.Analysis.Convex.Deriv import Mathlib.Analysis.Convex.EGauge diff --git a/Mathlib/Analysis/Convex/Continuous.lean b/Mathlib/Analysis/Convex/Continuous.lean new file mode 100644 index 0000000000000..746b4cf9313a6 --- /dev/null +++ b/Mathlib/Analysis/Convex/Continuous.lean @@ -0,0 +1,232 @@ +/- +Copyright (c) 2023 Yaël Dillies, Zichen Wang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Zichen Wang +-/ +import Mathlib.Analysis.Convex.Normed + +/-! +# Convex functions are continuous + +This file proves that a convex function from a finite dimensional real normed space to `ℝ` is +continuous. +-/ + +open FiniteDimensional Metric Set List Bornology +open scoped Topology + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + {C : Set E} {f : E → ℝ} {x₀ : E} {ε r r' M : ℝ} + +lemma ConvexOn.lipschitzOnWith_of_abs_le (hf : ConvexOn ℝ (ball x₀ r) f) (hε : 0 < ε) + (hM : ∀ a, dist a x₀ < r → |f a| ≤ M) : + LipschitzOnWith (2 * M / ε).toNNReal f (ball x₀ (r - ε)) := by + set K := 2 * M / ε with hK + have oneside {x y : E} (hx : x ∈ ball x₀ (r - ε)) (hy : y ∈ ball x₀ (r - ε)) : + f x - f y ≤ K * ‖x - y‖ := by + obtain rfl | hxy := eq_or_ne x y + · simp + have hx₀r : ball x₀ (r - ε) ⊆ ball x₀ r := ball_subset_ball <| by linarith + have hx' : x ∈ ball x₀ r := hx₀r hx + have hy' : y ∈ ball x₀ r := hx₀r hy + let z := x + (ε / ‖x - y‖) • (x - y) + replace hxy : 0 < ‖x - y‖ := by rwa [norm_sub_pos_iff] + have hz : z ∈ ball x₀ r := mem_ball_iff_norm.2 <| by + calc + _ = ‖(x - x₀) + (ε / ‖x - y‖) • (x - y)‖ := by simp only [z, add_sub_right_comm] + _ ≤ ‖x - x₀‖ + ‖(ε / ‖x - y‖) • (x - y)‖ := norm_add_le .. + _ < r - ε + ε := + add_lt_add_of_lt_of_le (mem_ball_iff_norm.1 hx) <| by + simp [norm_smul, abs_of_nonneg, hε.le, hxy.ne'] + _ = r := by simp + let a := ε / (ε + ‖x - y‖) + let b := ‖x - y‖ / (ε + ‖x - y‖) + have hab : a + b = 1 := by field_simp [a, b] + have hxyz : x = a • y + b • z := by + calc + x = a • x + b • x := by rw [Convex.combo_self hab] + _ = a • y + b • z := by simp [z, a, b, smul_smul, hxy.ne', smul_sub]; abel + rw [hK, mul_comm, ← mul_div_assoc, le_div_iff₀' hε] + calc + ε * (f x - f y) ≤ ‖x - y‖ * (f z - f x) := by + rw [mul_sub, mul_sub, sub_le_sub_iff, ← add_mul] + have h := hf.2 hy' hz (by positivity) (by positivity) hab + field_simp [← hxyz, a, b, ← mul_div_right_comm] at h + rwa [← le_div_iff₀' (by positivity), add_comm (_ * _)] + _ ≤ _ := by + rw [sub_eq_add_neg (f _), two_mul] + gcongr + · exact (le_abs_self _).trans <| hM _ hz + · exact (neg_le_abs _).trans <| hM _ hx' + refine .of_dist_le' fun x hx y hy ↦ ?_ + simp_rw [dist_eq_norm_sub, Real.norm_eq_abs, abs_sub_le_iff] + exact ⟨oneside hx hy, norm_sub_rev x _ ▸ oneside hy hx⟩ + +lemma ConcaveOn.lipschitzOnWith_of_abs_le (hf : ConcaveOn ℝ (ball x₀ r) f) (hε : 0 < ε) + (hM : ∀ a, dist a x₀ < r → |f a| ≤ M) : + LipschitzOnWith (2 * M / ε).toNNReal f (ball x₀ (r - ε)) := by + simpa using hf.neg.lipschitzOnWith_of_abs_le hε <| by simpa using hM + +lemma ConvexOn.exists_lipschitzOnWith_of_isBounded (hf : ConvexOn ℝ (ball x₀ r) f) (hr : r' < r) + (hf' : IsBounded (f '' ball x₀ r)) : ∃ K, LipschitzOnWith K f (ball x₀ r') := by + rw [isBounded_iff_subset_ball 0] at hf' + simp only [Set.subset_def, mem_image, mem_ball, dist_zero_right, Real.norm_eq_abs, + forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at hf' + obtain ⟨M, hM⟩ := hf' + rw [← sub_sub_cancel r r'] + exact ⟨_, hf.lipschitzOnWith_of_abs_le (sub_pos.2 hr) fun a ha ↦ (hM a ha).le⟩ + +lemma ConcaveOn.exists_lipschitzOnWith_of_isBounded (hf : ConcaveOn ℝ (ball x₀ r) f) (hr : r' < r) + (hf' : IsBounded (f '' ball x₀ r)) : ∃ K, LipschitzOnWith K f (ball x₀ r') := by + replace hf' : IsBounded ((-f) '' ball x₀ r) := by convert hf'.neg; ext; simp [neg_eq_iff_eq_neg] + simpa using hf.neg.exists_lipschitzOnWith_of_isBounded hr hf' + +lemma ConvexOn.isBoundedUnder_abs (hf : ConvexOn ℝ C f) {x₀ : E} (hC : C ∈ 𝓝 x₀) : + (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f| ↔ (𝓝 x₀).IsBoundedUnder (· ≤ ·) f := by + refine ⟨fun h ↦ h.mono_le <| .of_forall fun x ↦ le_abs_self _, ?_⟩ + rintro ⟨r, hr⟩ + refine ⟨|r| + 2 * |f x₀|, ?_⟩ + have : (𝓝 x₀).Tendsto (fun y => 2 • x₀ - y) (𝓝 x₀) := + tendsto_nhds_nhds.2 (⟨·, ·, by simp [two_nsmul, dist_comm]⟩) + simp only [Filter.eventually_map, Pi.abs_apply, abs_le'] at hr ⊢ + filter_upwards [this.eventually_mem hC, hC, hr, this.eventually hr] with y hx hx' hfr hfr' + refine ⟨hfr.trans <| (le_abs_self _).trans <| by simp, ?_⟩ + rw [← sub_le_iff_le_add, neg_sub_comm, sub_le_iff_le_add', ← abs_two, ← abs_mul] + calc + -|2 * f x₀| ≤ 2 * f x₀ := neg_abs_le _ + _ ≤ f y + f (2 • x₀ - y) := by + have := hf.2 hx' hx (by positivity) (by positivity) (add_halves _) + simp only [one_div, ← Nat.cast_smul_eq_nsmul ℝ, Nat.cast_ofNat, smul_sub, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, inv_smul_smul₀, add_sub_cancel, smul_eq_mul] at this + cancel_denoms at this + rwa [← Nat.cast_two, Nat.cast_smul_eq_nsmul] at this + _ ≤ f y + |r| := by gcongr; exact hfr'.trans (le_abs_self _) + +lemma ConcaveOn.isBoundedUnder_abs (hf : ConcaveOn ℝ C f) {x₀ : E} (hC : C ∈ 𝓝 x₀) : + (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f| ↔ (𝓝 x₀).IsBoundedUnder (· ≥ ·) f := by + simpa [Pi.neg_def, Pi.abs_def] using hf.neg.isBoundedUnder_abs hC + +lemma ConvexOn.continuousOn_tfae (hC : IsOpen C) (hC' : C.Nonempty) (hf : ConvexOn ℝ C f) : TFAE [ + LocallyLipschitzOn C f, + ContinuousOn f C, + ∃ x₀ ∈ C, ContinuousAt f x₀, + ∃ x₀ ∈ C, (𝓝 x₀).IsBoundedUnder (· ≤ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f|] := by + tfae_have 1 → 2 + · exact LocallyLipschitzOn.continuousOn + tfae_have 2 → 3 + · obtain ⟨x₀, hx₀⟩ := hC' + exact fun h ↦ ⟨x₀, hx₀, h.continuousAt <| hC.mem_nhds hx₀⟩ + tfae_have 3 → 4 + · rintro ⟨x₀, hx₀, h⟩ + exact ⟨x₀, hx₀, f x₀ + 1, by simpa using h.eventually (eventually_le_nhds (by simp))⟩ + tfae_have 4 → 5 + · rintro ⟨x₀, hx₀, r, hr⟩ x hx + have : ∀ᶠ δ in 𝓝 (0 : ℝ), (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀ ∈ C := by + have h : ContinuousAt (fun δ : ℝ ↦ (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀) 0 := by + fun_prop (disch := norm_num) + exact h (by simpa using hC.mem_nhds hx) + obtain ⟨δ, hδ₀, hy, hδ₁⟩ := (this.and <| eventually_lt_nhds zero_lt_one).exists_gt + set y := (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀ + refine ⟨max r (f y), ?_⟩ + simp only [Filter.eventually_map, Pi.abs_apply] at hr ⊢ + obtain ⟨ε, hε, hr⟩ := Metric.eventually_nhds_iff.1 <| hr.and (hC.eventually_mem hx₀) + refine Metric.eventually_nhds_iff.2 ⟨ε * δ, by positivity, fun z hz ↦ ?_⟩ + have hx₀' : δ⁻¹ • (x - y) + y = x₀ := MulAction.injective₀ (sub_ne_zero.2 hδ₁.ne') <| by + simp [y, smul_sub, smul_smul, hδ₀.ne', div_eq_mul_inv, sub_ne_zero.2 hδ₁.ne', mul_left_comm, + sub_mul, sub_smul] + let w := δ⁻¹ • (z - y) + y + have hwyz : δ • w + (1 - δ) • y = z := by simp [w, hδ₀.ne', sub_smul] + have hw : dist w x₀ < ε := by + simpa [w, ← hx₀', dist_smul₀, abs_of_nonneg, hδ₀.le, inv_mul_lt_iff', hδ₀] + calc + f z ≤ max (f w) (f y) := + hf.le_max_of_mem_segment (hr hw).2 hy ⟨_, _, hδ₀.le, sub_nonneg.2 hδ₁.le, by simp, hwyz⟩ + _ ≤ max r (f y) := by gcongr; exact (hr hw).1 + tfae_have 6 ↔ 5 + · exact forall₂_congr fun x₀ hx₀ ↦ hf.isBoundedUnder_abs (hC.mem_nhds hx₀) + tfae_have 6 → 1 + · rintro h x hx + obtain ⟨r, hr⟩ := h hx + obtain ⟨ε, hε, hεD⟩ := Metric.mem_nhds_iff.1 <| Filter.inter_mem (hC.mem_nhds hx) hr + simp only [preimage_setOf_eq, Pi.abs_apply, subset_inter_iff, hC.nhdsWithin_eq hx] at hεD ⊢ + obtain ⟨K, hK⟩ := exists_lipschitzOnWith_of_isBounded (hf.subset hεD.1 (convex_ball ..)) + (half_lt_self hε) <| isBounded_iff_forall_norm_le.2 ⟨r, by simpa using hεD.2⟩ + exact ⟨K, _, ball_mem_nhds _ (by simpa), hK⟩ + tfae_finish + +lemma ConcaveOn.continuousOn_tfae (hC : IsOpen C) (hC' : C.Nonempty) (hf : ConcaveOn ℝ C f) : TFAE [ + LocallyLipschitzOn C f, + ContinuousOn f C, + ∃ x₀ ∈ C, ContinuousAt f x₀, + ∃ x₀ ∈ C, (𝓝 x₀).IsBoundedUnder (· ≥ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≥ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f|] := by + have := hf.neg.continuousOn_tfae hC hC' + simp at this + convert this using 8 <;> exact (Equiv.neg ℝ).exists_congr (by simp) + +lemma ConvexOn.locallyLipschitzOn_iff_continuousOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn C f ↔ ContinuousOn f C := by + obtain rfl | hC' := C.eq_empty_or_nonempty + · simp + · exact (hf.continuousOn_tfae hC hC').out 0 1 + +lemma ConcaveOn.locallyLipschitzOn_iff_continuousOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn C f ↔ ContinuousOn f C := by + simpa using hf.neg.locallyLipschitzOn_iff_continuousOn hC + +variable [FiniteDimensional ℝ E] + +protected lemma ConvexOn.locallyLipschitzOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn C f := by + obtain rfl | ⟨x₀, hx₀⟩ := C.eq_empty_or_nonempty + · simp + · obtain ⟨b, hx₀b, hbC⟩ := exists_mem_interior_convexHull_affineBasis (hC.mem_nhds hx₀) + refine ((hf.continuousOn_tfae hC ⟨x₀, hx₀⟩).out 3 0).mp ?_ + refine ⟨x₀, hx₀, BddAbove.isBoundedUnder (IsOpen.mem_nhds isOpen_interior hx₀b) ?_⟩ + exact (hf.bddAbove_convexHull ((subset_convexHull ..).trans hbC) + ((finite_range _).image _).bddAbove).mono (by gcongr; exact interior_subset) + +protected lemma ConcaveOn.locallyLipschitzOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn C f := by simpa using hf.neg.locallyLipschitzOn hC + +protected lemma ConvexOn.continuousOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + ContinuousOn f C := (hf.locallyLipschitzOn hC).continuousOn + +protected lemma ConcaveOn.continuousOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + ContinuousOn f C := (hf.locallyLipschitzOn hC).continuousOn + +lemma ConvexOn.locallyLipschitzOn_interior (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn (interior C) f := + (hf.subset interior_subset hf.1.interior).locallyLipschitzOn isOpen_interior + +lemma ConcaveOn.locallyLipschitzOn_interior (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn (interior C) f := + (hf.subset interior_subset hf.1.interior).locallyLipschitzOn isOpen_interior + +lemma ConvexOn.continuousOn_interior (hf : ConvexOn ℝ C f) : ContinuousOn f (interior C) := + hf.locallyLipschitzOn_interior.continuousOn + +lemma ConcaveOn.continuousOn_interior (hf : ConcaveOn ℝ C f) : ContinuousOn f (interior C) := + hf.locallyLipschitzOn_interior.continuousOn + +protected lemma ConvexOn.locallyLipschitz (hf : ConvexOn ℝ univ f) : LocallyLipschitz f := by + simpa using hf.locallyLipschitzOn_interior + +protected lemma ConcaveOn.locallyLipschitz (hf : ConcaveOn ℝ univ f) : LocallyLipschitz f := by + simpa using hf.locallyLipschitzOn_interior + +-- Commented out since `intrinsicInterior` is not imported (but should be once these are proved) +-- proof_wanted ConvexOn.locallyLipschitzOn_intrinsicInterior (hf : ConvexOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConcaveOn.locallyLipschitzOn_intrinsicInterior (hf : ConcaveOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConvexOn.continuousOn_intrinsicInterior (hf : ConvexOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConcaveOn.continuousOn_intrinsicInterior (hf : ConcaveOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) diff --git a/Mathlib/Analysis/Convex/Intrinsic.lean b/Mathlib/Analysis/Convex/Intrinsic.lean index b244bd60d24fb..82cc1b1b0da2d 100644 --- a/Mathlib/Analysis/Convex/Intrinsic.lean +++ b/Mathlib/Analysis/Convex/Intrinsic.lean @@ -96,24 +96,6 @@ theorem intrinsicFrontier_subset_intrinsicClosure : intrinsicFrontier 𝕜 s ⊆ theorem subset_intrinsicClosure : s ⊆ intrinsicClosure 𝕜 s := fun x hx => ⟨⟨x, subset_affineSpan _ _ hx⟩, subset_closure hx, rfl⟩ -lemma intrinsicInterior_eq_interior_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicInterior 𝕜 s = interior s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' interior (f ⁻¹' s) = interior s - rw [f.image_interior, f.image_preimage] - -lemma intrinsicFrontier_eq_frontier_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicFrontier 𝕜 s = frontier s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' frontier (f ⁻¹' s) = frontier s - rw [f.image_frontier, f.image_preimage] - -lemma intrinsicClosure_eq_closure_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicClosure 𝕜 s = closure s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' closure (f ⁻¹' s) = closure s - rw [f.image_closure, f.image_preimage] - @[simp] theorem intrinsicInterior_empty : intrinsicInterior 𝕜 (∅ : Set P) = ∅ := by simp [intrinsicInterior] @@ -123,15 +105,6 @@ theorem intrinsicFrontier_empty : intrinsicFrontier 𝕜 (∅ : Set P) = ∅ := @[simp] theorem intrinsicClosure_empty : intrinsicClosure 𝕜 (∅ : Set P) = ∅ := by simp [intrinsicClosure] -@[simp] lemma intrinsicInterior_univ : intrinsicInterior 𝕜 (univ : Set P) = univ := by - simp [intrinsicInterior] - -@[simp] lemma intrinsicFrontier_univ : intrinsicFrontier 𝕜 (univ : Set P) = ∅ := by - simp [intrinsicFrontier] - -@[simp] lemma intrinsicClosure_univ : intrinsicClosure 𝕜 (univ : Set P) = univ := by - simp [intrinsicClosure] - @[simp] theorem intrinsicClosure_nonempty : (intrinsicClosure 𝕜 s).Nonempty ↔ s.Nonempty := ⟨by simp_rw [nonempty_iff_ne_empty]; rintro h rfl; exact h intrinsicClosure_empty, diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 69e02cd2a9319..b526794778cbb 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -609,7 +609,7 @@ theorem continuous_of_cover_nhds {ι : Sort*} {f : α → β} {s : ι → Set α rw [ContinuousAt, ← nhdsWithin_eq_nhds.2 hi] exact hf _ _ (mem_of_mem_nhds hi) -theorem continuousOn_empty (f : α → β) : ContinuousOn f ∅ := fun _ => False.elim +@[simp] theorem continuousOn_empty (f : α → β) : ContinuousOn f ∅ := fun _ => False.elim @[simp] theorem continuousOn_singleton (f : α → β) (a : α) : ContinuousOn f {a} := diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index f1113456868a1..4a3b5c75b8fd9 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -356,7 +356,9 @@ Single Variable Real Analysis: Weierstrass trigonometric approximation theorem: 'span_fourier_closure_eq_top' Convexity: convex functions of a real variable: 'ConvexOn' - continuity and differentiability of convex functions: 'https://en.wikipedia.org/wiki/Convex_function#Functions_of_one_variable' + continuity and differentiability of convex functions: + continuity: 'ConvexOn.continuousOn' + differentiability: 'https://en.wikipedia.org/wiki/Convex_function#Functions_of_one_variable' characterizations of convexity: 'convexOn_of_deriv2_nonneg' convexity inequalities: 'analysis/mean_inequalities.html' From 2cb73c7a2936f65cffe50de51138795da5df8939 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 11:49:14 +0000 Subject: [PATCH 080/174] refactor: Generalise the rearrangement inequality (#15815) --- Mathlib/Algebra/Order/Rearrangement.lean | 310 ++++++++++++----------- 1 file changed, 165 insertions(+), 145 deletions(-) diff --git a/Mathlib/Algebra/Order/Rearrangement.lean b/Mathlib/Algebra/Order/Rearrangement.lean index f032d3104e240..f69acca4d99d3 100644 --- a/Mathlib/Algebra/Order/Rearrangement.lean +++ b/Mathlib/Algebra/Order/Rearrangement.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys -/ import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Data.Prod.Lex @@ -41,23 +40,30 @@ convenience. The case for `Monotone`/`Antitone` pairs of functions over a `LinearOrder` is not deduced in this file because it is easily deducible from the `Monovary` API. + +## TODO + +Add equality cases for when the permute function is injective. This comes from the following fact: +If `Monovary f g`, `Injective g` and `σ` is a permutation, then `Monovary f (g ∘ σ) ↔ σ = 1`. -/ open Equiv Equiv.Perm Finset Function OrderDual -variable {ι α β : Type*} +variable {ι α β : Type*} [LinearOrderedSemiring α] [ExistsAddOfLE α] + [LinearOrderedCancelAddCommMonoid β] [Module α β] /-! ### Scalar multiplication versions -/ - section SMul -variable [LinearOrderedRing α] [LinearOrderedAddCommGroup β] [Module α β] [OrderedSMul α β] - {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} +/-! #### Weak rearrangement inequality -/ + +section weak_inequality +variable [PosSMulMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ +`f` and `g` monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_smul_comp_perm_le_sum_smul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) ≤ ∑ i ∈ s, f i • g i := by classical @@ -106,9 +112,62 @@ theorem MonovaryOn.sum_smul_comp_perm_le_sum_smul (hfg : MonovaryOn f g s) rintro rfl exact has hx.2 +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together on `s`. Stated by permuting the entries of `g`. -/ +theorem AntivaryOn.sum_smul_le_sum_smul_comp_perm (hfg : AntivaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f i • g (σ i) := + hfg.dual_right.sum_smul_comp_perm_le_sum_smul hσ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together on `s`. Stated by permuting the entries of `f`. -/ +theorem MonovaryOn.sum_comp_perm_smul_le_sum_smul (hfg : MonovaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i ≤ ∑ i ∈ s, f i • g i := by + convert hfg.sum_smul_comp_perm_le_sum_smul + (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 + exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together on `s`. Stated by permuting the entries of `f`. -/ +theorem AntivaryOn.sum_smul_le_sum_comp_perm_smul (hfg : AntivaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f (σ i) • g i := + hfg.dual_right.sum_comp_perm_smul_le_sum_smul hσ + +variable [Fintype ι] + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_smul_comp_perm_le_sum_smul (hfg : Monovary f g) : + ∑ i, f i • g (σ i) ≤ ∑ i, f i • g i := + (hfg.monovaryOn _).sum_smul_comp_perm_le_sum_smul fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ +theorem Antivary.sum_smul_le_sum_smul_comp_perm (hfg : Antivary f g) : + ∑ i, f i • g i ≤ ∑ i, f i • g (σ i) := + (hfg.antivaryOn _).sum_smul_le_sum_smul_comp_perm fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ +theorem Monovary.sum_comp_perm_smul_le_sum_smul (hfg : Monovary f g) : + ∑ i, f (σ i) • g i ≤ ∑ i, f i • g i := + (hfg.monovaryOn _).sum_comp_perm_smul_le_sum_smul fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ +theorem Antivary.sum_smul_le_sum_comp_perm_smul (hfg : Antivary f g) : + ∑ i, f i • g i ≤ ∑ i, f (σ i) • g i := + (hfg.antivaryOn _).sum_smul_le_sum_comp_perm_smul fun _ _ ↦ mem_univ _ + +end weak_inequality + +/-! #### Equality case of the rearrangement inequality -/ + +section equality_case +variable [PosSMulStrictMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} + /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ +`g`, which monovary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ MonovaryOn f (g ∘ σ) s := by @@ -134,26 +193,17 @@ theorem MonovaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : MonovaryOn f g s) · convert h.sum_smul_comp_perm_le_sum_smul ((set_support_inv_eq _).subset.trans hσ) using 1 simp_rw [Function.comp_apply, apply_inv_self] -/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ -theorem MonovaryOn.sum_smul_comp_perm_lt_sum_smul_iff (hfg : MonovaryOn f g s) +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +antivary together on `s`. Stated by permuting the entries of `g`. -/ +theorem AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := by - simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, - hfg.sum_smul_comp_perm_le_sum_smul hσ] - -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ -theorem MonovaryOn.sum_comp_perm_smul_le_sum_smul (hfg : MonovaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i ≤ ∑ i ∈ s, f i • g i := by - convert hfg.sum_smul_comp_perm_le_sum_smul - (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 - exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ + ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ AntivaryOn f (g ∘ σ) s := + (hfg.dual_right.sum_smul_comp_perm_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `f`. -/ +`g`, which monovary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ MonovaryOn (f ∘ σ) g s := by @@ -171,63 +221,93 @@ theorem MonovaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : MonovaryOn f g s) · rw [σ.symm.eq_preimage_iff_image_eq] exact Set.image_perm hσinv -/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not monovary together. Stated by permuting the entries of `f`. -/ -theorem MonovaryOn.sum_comp_perm_smul_lt_sum_smul_iff (hfg : MonovaryOn f g s) +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +antivary together on `s`. Stated by permuting the entries of `f`. -/ +theorem AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f (σ i) • g i < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn (f ∘ σ) g s := by - simp [← hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ, lt_iff_le_and_ne, - hfg.sum_comp_perm_smul_le_sum_smul hσ] + ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ AntivaryOn (f ∘ σ) g s := + (hfg.dual_right.sum_comp_perm_smul_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ -theorem AntivaryOn.sum_smul_le_sum_smul_comp_perm (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f i • g (σ i) := - hfg.dual_right.sum_smul_comp_perm_le_sum_smul hσ +@[deprecated (since := "2024-06-25")] +alias AntivaryOn.sum_smul_eq_sum_comp_perm_smul_iff := AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff + +variable [Fintype ι] + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary +together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Monovary f g) : + ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Monovary f (g ∘ σ) := by + simp [(hfg.monovaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary +together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Monovary f g) : + ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Monovary (f ∘ σ) g := by + simp [(hfg.monovaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary together. Stated by permuting the entries of `g`. -/ -theorem AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ AntivaryOn f (g ∘ σ) s := - (hfg.dual_right.sum_smul_comp_perm_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right +theorem Antivary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Antivary f g) : + ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Antivary f (g ∘ σ) := by + simp [(hfg.antivaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] @[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_smul_comp_perm_iff := AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff +alias Antivary.sum_smul_eq_sum_smul_comp_perm_iff := Antivary.sum_smul_comp_perm_eq_sum_smul_iff + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary +together. Stated by permuting the entries of `f`. -/ +theorem Antivary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Antivary f g) : + ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Antivary (f ∘ σ) g := by + simp [(hfg.antivaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] + +@[deprecated (since := "2024-06-25")] +alias Antivary.sum_smul_eq_sum_comp_perm_smul_iff := Antivary.sum_comp_perm_smul_eq_sum_smul_iff + +end equality_case + +/-! #### Strict rearrangement inequality -/ + +section strict_inequality +variable [PosSMulStrictMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not monovary together on `s`. Stated by permuting the entries of `g`. -/ +theorem MonovaryOn.sum_smul_comp_perm_lt_sum_smul_iff (hfg : MonovaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : + ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := by + simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, + hfg.sum_smul_comp_perm_le_sum_smul hσ] + +/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_smul_lt_sum_smul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i < ∑ i ∈ s, f i • g (σ i) ↔ ¬AntivaryOn f (g ∘ σ) s := by simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, eq_comm, hfg.sum_smul_le_sum_smul_comp_perm hσ] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ -theorem AntivaryOn.sum_smul_le_sum_comp_perm_smul (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f (σ i) • g i := by - convert hfg.sum_smul_le_sum_smul_comp_perm - (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 - exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ -theorem AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : AntivaryOn f g s) +/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not monovary together on `s`. Stated by permuting the entries of `f`. -/ +theorem MonovaryOn.sum_comp_perm_smul_lt_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ AntivaryOn (f ∘ σ) g s := - (hfg.dual_right.sum_comp_perm_smul_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right + ∑ i ∈ s, f (σ i) • g i < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn (f ∘ σ) g s := by + simp [← hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ, lt_iff_le_and_ne, + hfg.sum_comp_perm_smul_le_sum_smul hσ] @[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_comp_perm_smul_iff := AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff +alias AntivaryOn.sum_smul_eq_sum_smul_comp_perm_iff := AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_smul_lt_sum_comp_perm_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i < ∑ i ∈ s, f (σ i) • g i ↔ ¬AntivaryOn (f ∘ σ) g s := by @@ -236,19 +316,6 @@ theorem AntivaryOn.sum_smul_lt_sum_comp_perm_smul_iff (hfg : AntivaryOn f g s) variable [Fintype ι] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_smul_comp_perm_le_sum_smul (hfg : Monovary f g) : - ∑ i, f i • g (σ i) ≤ ∑ i, f i • g i := - (hfg.monovaryOn _).sum_smul_comp_perm_le_sum_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Monovary f g) : - ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Monovary f (g ∘ σ) := by - simp [(hfg.monovaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which monovary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ @@ -256,19 +323,6 @@ theorem Monovary.sum_smul_comp_perm_lt_sum_smul_iff (hfg : Monovary f g) : ∑ i, f i • g (σ i) < ∑ i, f i • g i ↔ ¬Monovary f (g ∘ σ) := by simp [(hfg.monovaryOn _).sum_smul_comp_perm_lt_sum_smul_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ -theorem Monovary.sum_comp_perm_smul_le_sum_smul (hfg : Monovary f g) : - ∑ i, f (σ i) • g i ≤ ∑ i, f i • g i := - (hfg.monovaryOn _).sum_comp_perm_smul_le_sum_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Monovary f g) : - ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Monovary (f ∘ σ) g := by - simp [(hfg.monovaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which monovary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ @@ -276,22 +330,6 @@ theorem Monovary.sum_comp_perm_smul_lt_sum_smul_iff (hfg : Monovary f g) : ∑ i, f (σ i) • g i < ∑ i, f i • g i ↔ ¬Monovary (f ∘ σ) g := by simp [(hfg.monovaryOn _).sum_comp_perm_smul_lt_sum_smul_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ -theorem Antivary.sum_smul_le_sum_smul_comp_perm (hfg : Antivary f g) : - ∑ i, f i • g i ≤ ∑ i, f i • g (σ i) := - (hfg.antivaryOn _).sum_smul_le_sum_smul_comp_perm fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary -together. Stated by permuting the entries of `g`. -/ -theorem Antivary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Antivary f g) : - ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Antivary f (g ∘ σ) := by - simp [(hfg.antivaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_smul_comp_perm_iff := Antivary.sum_smul_comp_perm_eq_sum_smul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ @@ -299,22 +337,6 @@ theorem Antivary.sum_smul_lt_sum_smul_comp_perm_iff (hfg : Antivary f g) : ∑ i, f i • g i < ∑ i, f i • g (σ i) ↔ ¬Antivary f (g ∘ σ) := by simp [(hfg.antivaryOn _).sum_smul_lt_sum_smul_comp_perm_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ -theorem Antivary.sum_smul_le_sum_comp_perm_smul (hfg : Antivary f g) : - ∑ i, f i • g i ≤ ∑ i, f (σ i) • g i := - (hfg.antivaryOn _).sum_smul_le_sum_comp_perm_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ -theorem Antivary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Antivary f g) : - ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Antivary (f ∘ σ) g := by - simp [(hfg.antivaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_comp_perm_smul_iff := Antivary.sum_comp_perm_smul_eq_sum_smul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is strictly decreased by a permutation if and only if `f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ @@ -322,6 +344,7 @@ theorem Antivary.sum_smul_lt_sum_comp_perm_smul_iff (hfg : Antivary f g) : ∑ i, f i • g i < ∑ i, f (σ i) • g i ↔ ¬Antivary (f ∘ σ) g := by simp [(hfg.antivaryOn _).sum_smul_lt_sum_comp_perm_smul_iff fun _ _ ↦ mem_univ _] +end strict_inequality end SMul /-! @@ -330,87 +353,84 @@ end SMul Special cases of the above when scalar multiplication is actually multiplication. -/ - section Mul - - -variable [LinearOrderedRing α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} +variable {s : Finset ι} {σ : Perm ι} {f g : ι → α} /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is maximized when `f` and -`g` monovary together. Stated by permuting the entries of `g`. -/ -theorem MonovaryOn.sum_mul_comp_perm_le_sum_mul (hfg : MonovaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) ≤ ∑ i ∈ s, f i * g i := +`g` monovary together on `s`. Stated by permuting the entries of `g`. -/ +theorem MonovaryOn.sum_mul_comp_perm_le_sum_mul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : + ∑ i ∈ s, f i * g (σ i) ≤ ∑ i ∈ s, f i * g i := hfg.sum_smul_comp_perm_le_sum_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ +which monovary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_mul_comp_perm_eq_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) = ∑ i ∈ s, f i * g i ↔ MonovaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_mul_comp_perm_lt_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_lt_sum_smul_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is maximized when `f` and -`g` monovary together. Stated by permuting the entries of `f`. -/ +`g` monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_le_sum_mul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i ≤ ∑ i ∈ s, f i * g i := hfg.sum_comp_perm_smul_le_sum_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `f`. -/ +which monovary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i = ∑ i ∈ s, f i * g i ↔ MonovaryOn (f ∘ σ) g s := hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not monovary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_lt_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i < ∑ i ∈ s, f i * g i ↔ ¬MonovaryOn (f ∘ σ) g s := hfg.sum_comp_perm_smul_lt_sum_smul_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is minimized when `f` and -`g` antivary together. Stated by permuting the entries of `g`. -/ +`g` antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_le_sum_mul_comp_perm (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i ≤ ∑ i ∈ s, f i * g (σ i) := hfg.sum_smul_le_sum_smul_comp_perm hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary -together. Stated by permuting the entries of `g`. -/ +which antivary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_eq_sum_mul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) = ∑ i ∈ s, f i * g i ↔ AntivaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_lt_sum_mul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i < ∑ i ∈ s, f i * g (σ i) ↔ ¬AntivaryOn f (g ∘ σ) s := hfg.sum_smul_lt_sum_smul_comp_perm_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is minimized when `f` and -`g` antivary together. Stated by permuting the entries of `f`. -/ +`g` antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_mul_le_sum_comp_perm_mul (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i ≤ ∑ i ∈ s, f (σ i) * g i := hfg.sum_smul_le_sum_comp_perm_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ +which antivary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i = ∑ i ∈ s, f i * g i ↔ AntivaryOn (f ∘ σ) g s := @@ -420,8 +440,8 @@ theorem AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : AntivaryOn f g s) alias AntivaryOn.sum_mul_eq_sum_comp_perm_mul_iff := AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_mul_lt_sum_comp_perm_mul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i < ∑ i ∈ s, f (σ i) * g i ↔ ¬AntivaryOn (f ∘ σ) g s := From 47ef55c55b51c82654295ce66768c1b0c7947575 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 14:00:34 +0000 Subject: [PATCH 081/174] refactor: Generalise special case of Jensen (#11448) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `Real.pow_sum_div_card_le_sum_pow` and `NNReal.pow_sum_div_card_le_sum_pow` both state that `(∑ i in s, f i) ^ (n + 1) / s.card ^ n ≤ ∑ i in s, f i ^ (n + 1)`, but one for `f : ι → ℝ` and the other for `f : ι → ℝ≥0`. Both lemmas can be unified by deriving the inequality for `f : ι → α` where `[LinearOrderedSemifield α] [ExistsAddOfLE α]`. This involves generalising the rearrangement inequality semirings. From LeanAPAP --- Mathlib/Algebra/Order/Chebyshev.lean | 75 +++++++++++++++-------- Mathlib/Analysis/MeanInequalitiesPow.lean | 22 ------- 2 files changed, 49 insertions(+), 48 deletions(-) diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index ebbe5643ee00c..439d79b0c4e2b 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -4,9 +4,12 @@ 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.Algebra.Order.Ring.Basic import Mathlib.GroupTheory.Perm.Cycle.Basic +import Mathlib.Tactic.GCongr +import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Positivity.Finset /-! # Chebyshev's sum inequality @@ -44,28 +47,26 @@ variable {ι α β : Type*} section SMul - -variable [LinearOrderedRing α] [LinearOrderedAddCommGroup β] [Module α β] [OrderedSMul α β] - {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} +variable [LinearOrderedSemiring α] [ExistsAddOfLE α] [LinearOrderedCancelAddCommMonoid β] + [Module α β] [OrderedSMul α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both 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.card • ∑ 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σ] - 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 _ + obtain ⟨σ, hσ, hs⟩ := s.countable_toSet.exists_cycleOn + rw [← card_range s.card, 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 _ /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the 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 := by - exact hfg.dual_right.sum_smul_sum_le_card_smul_sum + s.card • ∑ 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 ι] @@ -73,15 +74,15 @@ variable [Fintype ι] monotone/antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem Monovary.sum_smul_sum_le_card_smul_sum (hfg : Monovary f g) : - ((∑ i, f i) • ∑ i, g i) ≤ Fintype.card ι • ∑ i, f i • g i := + (∑ i, f i) • ∑ i, g i ≤ Fintype.card ι • ∑ i, f i • g i := (hfg.monovaryOn _).sum_smul_sum_le_card_smul_sum /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the other is antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem Antivary.card_smul_sum_le_sum_smul_sum (hfg : Antivary f g) : - (Fintype.card ι • ∑ i, f i • g i) ≤ (∑ i, f i) • ∑ i, g i := by - exact (hfg.dual_right.monovaryOn _).sum_smul_sum_le_card_smul_sum + Fintype.card ι • ∑ i, f i • g i ≤ (∑ i, f i) • ∑ i, g i := + (hfg.dual_right.monovaryOn _).sum_smul_sum_le_card_smul_sum end SMul @@ -93,14 +94,13 @@ Special cases of the above when scalar multiplication is actually multiplication section Mul - -variable [LinearOrderedRing α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} +variable [LinearOrderedSemiring α] [ExistsAddOfLE α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} /-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both 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.card * ∑ i ∈ s, f i * g i := by rw [← nsmul_eq_mul] exact hfg.sum_smul_sum_le_card_smul_sum @@ -108,10 +108,26 @@ 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.card : α) * ∑ 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) + | 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 + 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 + 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 @@ -124,25 +140,32 @@ variable [Fintype ι] monotone/antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem Monovary.sum_mul_sum_le_card_mul_sum (hfg : Monovary f g) : - ((∑ i, f i) * ∑ i, g i) ≤ Fintype.card ι * ∑ i, f i * g i := + (∑ i, f i) * ∑ i, g i ≤ Fintype.card ι * ∑ i, f i * g i := (hfg.monovaryOn _).sum_mul_sum_le_card_mul_sum /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the other is antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem Antivary.card_mul_sum_le_sum_mul_sum (hfg : Antivary f g) : - ((Fintype.card ι : α) * ∑ i, f i * g i) ≤ (∑ i, f i) * ∑ i, g i := + Fintype.card ι * ∑ i, f i * g i ≤ (∑ i, f i) * ∑ i, g i := (hfg.antivaryOn _).card_mul_sum_le_sum_mul_sum end Mul -variable [LinearOrderedField α] {s : Finset ι} {f : ι → α} +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 + 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 obtain rfl | hs := s.eq_empty_or_nonempty · simp - rw [← card_pos, ← @Nat.cast_pos α] at hs - rw [div_pow, div_le_div_iff (sq_pos_of_ne_zero hs.ne') hs, sq (s.card : α), mul_left_comm, ← - mul_assoc] - exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq hs.le + rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (s.card : α), 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/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index daa489d39347c..fdee6d8fdbe46 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -61,23 +61,6 @@ theorem pow_arith_mean_le_arith_mean_pow_of_even (w z : ι → ℝ) (hw : ∀ i (∑ i ∈ s, w i * z i) ^ n ≤ ∑ i ∈ s, w i * z i ^ n := hn.convexOn_pow.map_sum_le hw hw' fun _ _ => Set.mem_univ _ -/-- Specific case of Jensen's inequality for sums of powers -/ -theorem pow_sum_div_card_le_sum_pow {f : ι → ℝ} (n : ℕ) (hf : ∀ a ∈ s, 0 ≤ f a) : - (∑ x ∈ s, f x) ^ (n + 1) / (s.card : ℝ) ^ n ≤ ∑ x ∈ s, f x ^ (n + 1) := by - rcases s.eq_empty_or_nonempty with (rfl | hs) - · simp_rw [Finset.sum_empty, zero_pow n.succ_ne_zero, zero_div]; rfl - · have hs0 : 0 < (s.card : ℝ) := Nat.cast_pos.2 hs.card_pos - suffices (∑ x ∈ s, f x / s.card) ^ (n + 1) ≤ ∑ x ∈ s, f x ^ (n + 1) / s.card by - rwa [← Finset.sum_div, ← Finset.sum_div, div_pow, pow_succ (s.card : ℝ), ← div_div, - div_le_iff₀ hs0, div_mul, div_self hs0.ne', div_one] at this - have := - @ConvexOn.map_sum_le ℝ ℝ ℝ ι _ _ _ _ _ _ (Set.Ici 0) (fun x => x ^ (n + 1)) s - (fun _ => 1 / s.card) ((↑) ∘ f) (convexOn_pow (n + 1)) ?_ ?_ fun i hi => - Set.mem_Ici.2 (hf i hi) - · simpa only [inv_mul_eq_div, one_div, Algebra.id.smul_eq_mul] using this - · simp only [one_div, inv_nonneg, Nat.cast_nonneg, imp_true_iff] - · simpa only [one_div, Finset.sum_const, nsmul_eq_mul] using mul_inv_cancel₀ hs0.ne' - theorem zpow_arith_mean_le_arith_mean_zpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 < z i) (m : ℤ) : (∑ i ∈ s, w i * z i) ^ m ≤ ∑ i ∈ s, w i * z i ^ m := @@ -111,11 +94,6 @@ theorem pow_arith_mean_le_arith_mean_pow (w z : ι → ℝ≥0) (hw' : ∑ i ∈ Real.pow_arith_mean_le_arith_mean_pow s _ _ (fun i _ => (w i).coe_nonneg) (mod_cast hw') (fun i _ => (z i).coe_nonneg) n -theorem pow_sum_div_card_le_sum_pow (f : ι → ℝ≥0) (n : ℕ) : - (∑ x ∈ s, f x) ^ (n + 1) / (s.card : ℝ) ^ n ≤ ∑ x ∈ s, f x ^ (n + 1) := by - simpa only [← NNReal.coe_le_coe, NNReal.coe_sum, Nonneg.coe_div, NNReal.coe_pow] using - @Real.pow_sum_div_card_le_sum_pow ι s (((↑) : ℝ≥0 → ℝ) ∘ f) n fun _ _ => NNReal.coe_nonneg _ - /-- Weighted generalized mean inequality, version for sums over finite sets, with `ℝ≥0`-valued functions and real exponents. -/ theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ≥0) (hw' : ∑ i ∈ s, w i = 1) {p : ℝ} From 26ad5ceac5af421995494f631a2c55c3277826e5 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 1 Oct 2024 14:00:36 +0000 Subject: [PATCH 082/174] feat(RingTheory/TwoSidedIdeal): add some basic operations on two-sided-ideals (#14460) Co-authored-by: Jireh Loreaux Co-authored-by: zjj Co-authored-by: Jireh Loreaux --- Mathlib.lean | 2 + Mathlib/RingTheory/Congruence/Basic.lean | 10 + .../TwoSidedIdeal/BigOperators.lean | 74 +++++ .../RingTheory/TwoSidedIdeal/Operations.lean | 272 ++++++++++++++++++ 4 files changed, 358 insertions(+) create mode 100644 Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean create mode 100644 Mathlib/RingTheory/TwoSidedIdeal/Operations.lean diff --git a/Mathlib.lean b/Mathlib.lean index b7d0a3ed8d92f..4b4ebd9d3c7a4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4069,8 +4069,10 @@ import Mathlib.RingTheory.TensorProduct.MvPolynomial import Mathlib.RingTheory.Trace.Basic import Mathlib.RingTheory.Trace.Defs import Mathlib.RingTheory.TwoSidedIdeal.Basic +import Mathlib.RingTheory.TwoSidedIdeal.BigOperators import Mathlib.RingTheory.TwoSidedIdeal.Instances import Mathlib.RingTheory.TwoSidedIdeal.Lattice +import Mathlib.RingTheory.TwoSidedIdeal.Operations import Mathlib.RingTheory.UniqueFactorizationDomain import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Unramified.Derivations diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index 1a9c8ef824451..c6c46c89d1b9b 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -126,6 +126,16 @@ theorem ext' {c d : RingCon R} (H : ⇑c = ⇑d) : c = d := DFunLike.coe_injecti theorem ext {c d : RingCon R} (H : ∀ x y, c x y ↔ d x y) : c = d := ext' <| by ext; apply H +/-- +Pulling back a `RingCon` across a ring homomorphism. +-/ +def comap {R R' F : Type*} [Add R] [Add R'] + [FunLike F R R'] [AddHomClass F R R'] [Mul R] [Mul R'] [MulHomClass F R R'] + (J : RingCon R') (f : F) : + RingCon R where + __ := J.toCon.comap f (map_mul f) + __ := J.toAddCon.comap f (map_add f) + end Basic section Quotient diff --git a/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean b/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean new file mode 100644 index 0000000000000..8ed5a222e9ed2 --- /dev/null +++ b/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ + +import Mathlib.RingTheory.Congruence.BigOperators +import Mathlib.RingTheory.TwoSidedIdeal.Basic + +/-! +# Interactions between `∑, ∏` and two sided ideals + +-/ + +namespace TwoSidedIdeal + +section sum + +variable {R : Type*} [NonUnitalNonAssocRing R] (I : TwoSidedIdeal R) + +lemma listSum_mem {ι : Type*} (l : List ι) (f : ι → R) (hl : ∀ x ∈ l, f x ∈ I) : + (l.map f).sum ∈ I := by + rw [mem_iff, ← List.sum_map_zero] + exact I.ringCon.listSum l hl + +lemma multisetSum_mem {ι : Type*} (s : Multiset ι) (f : ι → R) (hs : ∀ x ∈ s, f x ∈ I) : + (s.map f).sum ∈ I := by + rw [mem_iff, ← Multiset.sum_map_zero] + exact I.ringCon.multisetSum s hs + +lemma finsetSum_mem {ι : Type*} (s : Finset ι) (f : ι → R) (hs : ∀ x ∈ s, f x ∈ I) : + s.sum f ∈ I := by + rw [mem_iff, ← Finset.sum_const_zero] + exact I.ringCon.finsetSum s hs + +end sum + +section prod + +section ring + +variable {R : Type*} [Ring R] (I : TwoSidedIdeal R) + +lemma listProd_mem {ι : Type*} (l : List ι) (f : ι → R) (hl : ∃ x ∈ l, f x ∈ I) : + (l.map f).prod ∈ I := by + induction l with + | nil => simp only [List.not_mem_nil, false_and, exists_false] at hl + | cons x l ih => + simp only [List.mem_cons, exists_eq_or_imp] at hl + rcases hl with h | hal + · simpa only [List.map_cons, List.prod_cons] using I.mul_mem_right _ _ h + · simpa only [List.map_cons, List.prod_cons] using I.mul_mem_left _ _ <| ih hal + +end ring + +section commRing + +variable {R : Type*} [CommRing R] (I : TwoSidedIdeal R) + +lemma multiSetProd_mem {ι : Type*} (s : Multiset ι) (f : ι → R) (hs : ∃ x ∈ s, f x ∈ I) : + (s.map f).prod ∈ I := by + rcases s + simpa using listProd_mem (hl := hs) + +lemma finsetProd_mem {ι : Type*} (s : Finset ι) (f : ι → R) (hs : ∃ x ∈ s, f x ∈ I) : + s.prod f ∈ I := by + rcases s + simpa using multiSetProd_mem (hs := hs) + +end commRing + +end prod + +end TwoSidedIdeal diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean new file mode 100644 index 0000000000000..ea08a0c3e8866 --- /dev/null +++ b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean @@ -0,0 +1,272 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang, Jireh Loreaux +-/ + +import Mathlib.RingTheory.TwoSidedIdeal.Lattice +import Mathlib.RingTheory.Congruence.Opposite +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Data.Fintype.BigOperators +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.Order.GaloisConnection + +/-! +# Operations on two-sided ideals + +This file defines operations on two-sided ideals of a ring `R`. + +## Main definitions and results + +- `TwoSidedIdeal.span`: the span of `s ⊆ R` is the smallest two-sided ideal containing the set. +- `TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure_nonunital`: in an associative but non-unital + ring, an element `x` is in the two-sided ideal spanned by `s` if and only if `x` is in the closure + of `s ∪ {y * a | y ∈ s, a ∈ R} ∪ {a * y | y ∈ s, a ∈ R} ∪ {a * y * b | y ∈ s, a, b ∈ R}` as an + additive subgroup. +- `TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure`: in a unital and associative ring, an + element `x` is in the two-sided ideal spanned by `s` if and only if `x` is in the closure of + `{a*y*b | a, b ∈ R, y ∈ s}` as an additive subgroup. + + +- `TwoSidedIdeal.comap`: pullback of a two-sided ideal; defined as the preimage of a + two-sided ideal. +- `TwoSidedIdeal.map`: pushforward of a two-sided ideal; defined as the span of the image of a + two-sided ideal. +- `TwoSidedIdeal.ker`: the kernel of a ring homomorphism as a two-sided ideal. + +- `TwoSidedIdeal.gc`: `fromIdeal` and `asIdeal` form a Galois connection where + `fromIdeal : Ideal R → TwoSidedIdeal R` is defined as the smallest two-sided ideal containing an + ideal and `asIdeal : TwoSidedIdeal R → Ideal R` the inclusion map. +-/ + +namespace TwoSidedIdeal + +section NonUnitalNonAssocRing + +variable {R S : Type*} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] +variable {F : Type*} [FunLike F R S] +variable (f : F) + +/-- +The smallest two-sided ideal containing a set. +-/ +abbrev span (s : Set R) : TwoSidedIdeal R := + { ringCon := ringConGen (fun a b ↦ a - b ∈ s) } + +lemma subset_span {s : Set R} : s ⊆ (span s : Set R) := by + intro x hx + rw [SetLike.mem_coe, mem_iff] + exact RingConGen.Rel.of _ _ (by simpa using hx) + +lemma mem_span_iff {s : Set R} {x} : + x ∈ span s ↔ ∀ (I : TwoSidedIdeal R), s ⊆ I → x ∈ I := by + refine ⟨?_, fun h => h _ subset_span⟩ + delta span + rw [RingCon.ringConGen_eq] + intro h I hI + refine sInf_le (α := RingCon R) ?_ h + intro x y hxy + specialize hI hxy + rwa [SetLike.mem_coe, ← rel_iff] at hI + +lemma span_mono {s t : Set R} (h : s ⊆ t) : span s ≤ span t := by + intro x hx + rw [mem_span_iff] at hx ⊢ + exact fun I hI => hx I <| h.trans hI + +/-- +Pushout of a two-sided ideal. Defined as the span of the image of a two-sided ideal under a ring +homomorphism. +-/ +def map (I : TwoSidedIdeal R) : TwoSidedIdeal S := + span (f '' I) + +lemma map_mono {I J : TwoSidedIdeal R} (h : I ≤ J) : + map f I ≤ map f J := + span_mono <| Set.image_mono h + +variable [NonUnitalRingHomClass F R S] + +/-- +Preimage of a two-sided ideal, as a two-sided ideal. -/ +def comap (I : TwoSidedIdeal S) : TwoSidedIdeal R := +{ ringCon := I.ringCon.comap f } + +lemma mem_comap {I : TwoSidedIdeal S} {x : R} : + x ∈ I.comap f ↔ f x ∈ I := by + simp [comap, RingCon.comap, mem_iff] + +/-- +The kernel of a ring homomorphism, as a two-sided ideal. +-/ +def ker : TwoSidedIdeal R := + .mk' + {r | f r = 0} (map_zero _) (by rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2]) + (by rintro _ (h : f _ = 0); simp [h]) (by rintro _ _ (h : f _ = 0); simp [h]) + (by rintro _ _ (h : f _ = 0); simp [h]) + +lemma mem_ker {x : R} : x ∈ ker f ↔ f x = 0 := by + delta ker; rw [mem_mk']; rfl + · rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2] + · rintro _ (h : f _ = 0); simp [h] + · rintro _ _ (h : f _ = 0); simp [h] + · rintro _ _ (h : f _ = 0); simp [h] + +end NonUnitalNonAssocRing + +section NonUnitalRing + +variable {R : Type*} [NonUnitalRing R] + +open AddSubgroup in +/-- If `s : Set R` is absorbing under multiplication, then its `TwoSidedIdeal.span` coincides with +its `AddSubgroup.closure`, as sets. -/ +lemma mem_span_iff_mem_addSubgroup_closure_absorbing {s : Set R} + (h_left : ∀ x y, y ∈ s → x * y ∈ s) (h_right : ∀ y x, y ∈ s → y * x ∈ s) {z : R} : + z ∈ span s ↔ z ∈ closure s := by + have h_left' {x y} (hy : y ∈ closure s) : x * y ∈ closure s := by + have := (AddMonoidHom.mulLeft x).map_closure s ▸ mem_map_of_mem _ hy + refine closure_mono ?_ this + rintro - ⟨y, hy, rfl⟩ + exact h_left x y hy + have h_right' {y x} (hy : y ∈ closure s) : y * x ∈ closure s := by + have := (AddMonoidHom.mulRight x).map_closure s ▸ mem_map_of_mem _ hy + refine closure_mono ?_ this + rintro - ⟨y, hy, rfl⟩ + exact h_right y x hy + let I : TwoSidedIdeal R := .mk' (closure s) (AddSubgroup.zero_mem _) + (AddSubgroup.add_mem _) (AddSubgroup.neg_mem _) h_left' h_right' + suffices z ∈ span s ↔ z ∈ I by simpa only [I, mem_mk', SetLike.mem_coe] + rw [mem_span_iff] + -- Suppose that for every ideal `J` with `s ⊆ J`, then `z ∈ J`. Apply this to `I` to get `z ∈ I`. + refine ⟨fun h ↦ h I fun x hx ↦ ?mem_closure_of_forall, fun hz J hJ ↦ ?mem_ideal_of_subset⟩ + case mem_closure_of_forall => simpa only [I, SetLike.mem_coe, mem_mk'] using subset_closure hx + /- Conversely, suppose that `z ∈ I` and that `J` is any ideal containing `s`. Then by the + induction principle for `AddSubgroup`, we must also have `z ∈ J`. -/ + case mem_ideal_of_subset => + simp only [I, SetLike.mem_coe, mem_mk'] at hz + induction hz using closure_induction' with + | mem x hx => exact hJ hx + | one => exact zero_mem _ + | mul x _ y _ hx hy => exact J.add_mem hx hy + | inv x _ hx => exact J.neg_mem hx + +open Pointwise Set + +lemma set_mul_subset {s : Set R} {I : TwoSidedIdeal R} (h : s ⊆ I) (t : Set R): + t * s ⊆ I := by + rintro - ⟨r, -, x, hx, rfl⟩ + exact mul_mem_left _ _ _ (h hx) + +lemma subset_mul_set {s : Set R} {I : TwoSidedIdeal R} (h : s ⊆ I) (t : Set R): + s * t ⊆ I := by + rintro - ⟨x, hx, r, -, rfl⟩ + exact mul_mem_right _ _ _ (h hx) + +lemma mem_span_iff_mem_addSubgroup_closure_nonunital {s : Set R} {z : R} : + z ∈ span s ↔ z ∈ AddSubgroup.closure (s ∪ s * univ ∪ univ * s ∪ univ * s * univ) := by + trans z ∈ span (s ∪ s * univ ∪ univ * s ∪ univ * s * univ) + · refine ⟨(span_mono (by simp only [Set.union_assoc, Set.subset_union_left]) ·), fun h ↦ ?_⟩ + refine mem_span_iff.mp h (span s) ?_ + simp only [union_subset_iff, union_assoc] + exact ⟨subset_span, subset_mul_set subset_span _, set_mul_subset subset_span _, + subset_mul_set (set_mul_subset subset_span _) _⟩ + · refine mem_span_iff_mem_addSubgroup_closure_absorbing ?_ ?_ + · rintro x y (((hy | ⟨y, hy, r, -, rfl⟩) | ⟨r, -, y, hy, rfl⟩) | + ⟨-, ⟨r', -, y, hy, rfl⟩, r, -, rfl⟩) + · exact .inl <| .inr <| ⟨x, mem_univ _, y, hy, rfl⟩ + · exact .inr <| ⟨x * y, ⟨x, mem_univ _, y, hy, rfl⟩, r, mem_univ _, mul_assoc ..⟩ + · exact .inl <| .inr <| ⟨x * r, mem_univ _, y, hy, mul_assoc ..⟩ + · refine .inr <| ⟨x * r' * y, ⟨x * r', mem_univ _, y, hy, ?_⟩, ⟨r, mem_univ _, ?_⟩⟩ + all_goals simp [mul_assoc] + · rintro y x (((hy | ⟨y, hy, r, -, rfl⟩) | ⟨r, -, y, hy, rfl⟩) | + ⟨-, ⟨r', -, y, hy, rfl⟩, r, -, rfl⟩) + · exact .inl <| .inl <| .inr ⟨y, hy, x, mem_univ _, rfl⟩ + · exact .inl <| .inl <| .inr ⟨y, hy, r * x, mem_univ _, (mul_assoc ..).symm⟩ + · exact .inr <| ⟨r * y, ⟨r, mem_univ _, y, hy, rfl⟩, x, mem_univ _, rfl⟩ + · refine .inr <| ⟨r' * y, ⟨r', mem_univ _, y, hy, rfl⟩, r * x, mem_univ _, ?_⟩ + simp [mul_assoc] + +end NonUnitalRing + +section Ring + +variable {R : Type*} [Ring R] + +open Pointwise Set in +lemma mem_span_iff_mem_addSubgroup_closure {s : Set R} {z : R} : + z ∈ span s ↔ z ∈ AddSubgroup.closure (univ * s * univ) := by + trans z ∈ span (univ * s * univ) + · refine ⟨(span_mono (fun x hx ↦ ?_) ·), fun hz ↦ ?_⟩ + · exact ⟨1 * x, ⟨1, mem_univ _, x, hx, rfl⟩, 1, mem_univ _, by simp⟩ + · exact mem_span_iff.mp hz (span s) <| subset_mul_set (set_mul_subset subset_span _) _ + · refine mem_span_iff_mem_addSubgroup_closure_absorbing ?_ ?_ + · intro x y hy + rw [mul_assoc] at hy ⊢ + obtain ⟨r, -, y, hy, rfl⟩ := hy + exact ⟨x * r, mem_univ _, y, hy, mul_assoc ..⟩ + · rintro - x ⟨y, hy, r, -, rfl⟩ + exact ⟨y, hy, r * x, mem_univ _, (mul_assoc ..).symm⟩ + +/-- Given an ideal `I`, `span I` is the smallest two-sided ideal containing `I`. -/ +def fromIdeal : Ideal R →o TwoSidedIdeal R where + toFun I := span I + monotone' _ _ := span_mono + +lemma mem_fromIdeal {I : Ideal R} {x : R} : + x ∈ fromIdeal I ↔ x ∈ span I := by simp [fromIdeal] + +/-- Every two-sided ideal is also a left ideal. -/ +def asIdeal : TwoSidedIdeal R →o Ideal R where + toFun I := + { carrier := I + add_mem' := I.add_mem + zero_mem' := I.zero_mem + smul_mem' := fun r x hx => I.mul_mem_left r x hx } + monotone' _ _ h _ h' := h h' + +lemma mem_asIdeal {I : TwoSidedIdeal R} {x : R} : + x ∈ asIdeal I ↔ x ∈ I := by simp [asIdeal] + +lemma gc : GaloisConnection fromIdeal (asIdeal (R := R)) := + fun I J => ⟨fun h x hx ↦ h <| mem_span_iff.2 fun _ H ↦ H hx, fun h x hx ↦ by + simp only [fromIdeal, OrderHom.coe_mk, mem_span_iff] at hx + exact hx _ h⟩ + +@[simp] +lemma coe_asIdeal {I : TwoSidedIdeal R} : (asIdeal I : Set R) = I := rfl + +/-- Every two-sided ideal is also a right ideal. -/ +def asIdealOpposite : TwoSidedIdeal R →o Ideal Rᵐᵒᵖ where + toFun I := asIdeal ⟨I.ringCon.op⟩ + monotone' I J h x h' := by + simp only [mem_asIdeal, mem_iff, RingCon.op_iff, MulOpposite.unop_zero] at h' ⊢ + exact J.rel_iff _ _ |>.2 <| h <| I.rel_iff 0 x.unop |>.1 h' + +lemma mem_asIdealOpposite {I : TwoSidedIdeal R} {x : Rᵐᵒᵖ} : + x ∈ asIdealOpposite I ↔ x.unop ∈ I := by + simpa [asIdealOpposite, asIdeal, TwoSidedIdeal.mem_iff, RingCon.op_iff] using + ⟨I.ringCon.symm, I.ringCon.symm⟩ + +end Ring + +section CommRing + +variable {R : Type*} [CommRing R] + +/-- +When the ring is commutative, two-sided ideals are exactly the same as left ideals. +-/ +def orderIsoIdeal : TwoSidedIdeal R ≃o Ideal R where + toFun := asIdeal + invFun := fromIdeal + map_rel_iff' := ⟨fun h _ hx ↦ h hx, fun h ↦ asIdeal.monotone' h⟩ + left_inv _ := SetLike.ext fun _ ↦ mem_span_iff.trans <| by aesop + right_inv J := SetLike.ext fun x ↦ mem_span_iff.trans + ⟨fun h ↦ mem_mk' _ _ _ _ _ _ _ |>.1 <| h (mk' + J J.zero_mem J.add_mem J.neg_mem (J.mul_mem_left _) (J.mul_mem_right _)) + (fun x => by simp [mem_mk']), by aesop⟩ + +end CommRing + +end TwoSidedIdeal From 608231ac44965a572dac873dc6bfbc7e2471eedc Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 1 Oct 2024 14:00:37 +0000 Subject: [PATCH 083/174] feat(NumberTheory/JacobiSum/Basic): add lemmas on value of Jacobi sum (#17003) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds statements that the Jacobi sum of two characters of order dividing `n` is in `ℤ[μ]`, where `μ` is a primitive `n`th root of unity, and more precisely, is of the form `-1 + z*(μ-1)^2` with `z` in `ℤ[μ]`. Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- Mathlib/NumberTheory/JacobiSum/Basic.lean | 72 +++++++++++++++++++++++ 1 file changed, 72 insertions(+) diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index 5025b602f0542..1859c852cef47 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.NumberTheory.GaussSum +import Mathlib.NumberTheory.MulChar.Lemmas +import Mathlib.RingTheory.RootsOfUnity.Lemmas /-! # Jacobi Sums @@ -226,3 +228,73 @@ lemma jacobiSum_mul_jacobiSum_inv (h : ringChar F' ≠ ringChar F) {χ φ : MulC ← mul_inv, gaussSum_mul_gaussSum_eq_card Hχφ ψ.prim] end field_field + +section image + +variable {F R : Type*} [Fintype F] [Field F] [CommRing R] [IsDomain R] + +/-- If `χ` and `φ` are multiplicative characters on a finite field `F` satisfying `χ^n = φ^n = 1` +and with values in an integral domain `R`, and `μ` is a primitive `n`th root of unity in `R`, +then the Jacobi sum `J(χ,φ)` is in `ℤ[μ] ⊆ R`. -/ +lemma jacobiSum_mem_algebraAdjoin_of_pow_eq_one {n : ℕ} (hn : n ≠ 0) {χ φ : MulChar F R} + (hχ : χ ^ n = 1) (hφ : φ ^ n = 1) {μ : R} (hμ : IsPrimitiveRoot μ n) : + jacobiSum χ φ ∈ Algebra.adjoin ℤ {μ} := + Subalgebra.sum_mem _ fun _ _ ↦ Subalgebra.mul_mem _ + (MulChar.apply_mem_algebraAdjoin_of_pow_eq_one hn hχ hμ _) + (MulChar.apply_mem_algebraAdjoin_of_pow_eq_one hn hφ hμ _) + +open Algebra in +private +lemma MulChar.exists_apply_sub_one_eq_mul_sub_one {n : ℕ} (hn : n ≠ 0) {χ : MulChar F R} {μ : R} + (hχ : χ ^ n = 1) (hμ : IsPrimitiveRoot μ n) {x : F} (hx : x ≠ 0) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, χ x - 1 = z * (μ - 1) := by + obtain ⟨k, _, hk⟩ := exists_apply_eq_pow hn hχ hμ hx + refine hk ▸ ⟨(Finset.range k).sum (μ ^ ·), ?_, (geom_sum_mul μ k).symm⟩ + exact Subalgebra.sum_mem _ fun m _ ↦ Subalgebra.pow_mem _ (self_mem_adjoin_singleton _ μ) _ + +private +lemma MulChar.exists_apply_sub_one_mul_apply_sub_one {n : ℕ} (hn : n ≠ 0) {χ ψ : MulChar F R} + {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hμ : IsPrimitiveRoot μ n) (x : F) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, (χ x - 1) * (ψ (1 - x) - 1) = z * (μ - 1) ^ 2 := by + rcases eq_or_ne x 0 with rfl | hx₀ + · exact ⟨0, Subalgebra.zero_mem _, by rw [sub_zero, ψ.map_one, sub_self, mul_zero, zero_mul]⟩ + rcases eq_or_ne x 1 with rfl | hx₁ + · exact ⟨0, Subalgebra.zero_mem _, by rw [χ.map_one, sub_self, zero_mul, zero_mul]⟩ + obtain ⟨z₁, hz₁, Hz₁⟩ := MulChar.exists_apply_sub_one_eq_mul_sub_one hn hχ hμ hx₀ + obtain ⟨z₂, hz₂, Hz₂⟩ := + MulChar.exists_apply_sub_one_eq_mul_sub_one hn hψ hμ (sub_ne_zero_of_ne hx₁.symm) + rewrite [Hz₁, Hz₂, sq] + exact ⟨z₁ * z₂, Subalgebra.mul_mem _ hz₁ hz₂, mul_mul_mul_comm ..⟩ + +/-- If `χ` and `ψ` are multiplicative characters of order dividing `n` on a finite field `F` +with values in an integral domain `R` and `μ` is a primitive `n`th root of unity in `R`, +then `J(χ,ψ) = -1 + z*(μ - 1)^2` for some `z ∈ ℤ[μ] ⊆ R`. (We assume that `#F ≡ 1 mod n`.) +Note that we do not state this as a divisbility in `R`, as this would give a weaker statement. -/ +lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} + {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hn' : n ∣ Fintype.card F - 1) + (hμ : IsPrimitiveRoot μ n) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, jacobiSum χ ψ = -1 + z * (μ - 1) ^ 2 := by + obtain ⟨q, hq⟩ := hn' + rw [Nat.sub_eq_iff_eq_add NeZero.one_le] at hq + obtain ⟨z₁, hz₁, Hz₁⟩ := hμ.self_sub_one_pow_dvd_order hn + by_cases hχ₀ : χ = 1 <;> by_cases hψ₀ : ψ = 1 + · rw [hχ₀, hψ₀, jacobiSum_one_one] + refine ⟨q * z₁, Subalgebra.mul_mem _ (Subalgebra.natCast_mem _ q) hz₁, ?_⟩ + rw [hq, Nat.cast_add, Nat.cast_mul, Hz₁] + ring + · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ + rw [hχ₀, jacobiSum_one_nontrivial hψ₀, zero_mul, add_zero] + · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ + rw [jacobiSum_comm, hψ₀, jacobiSum_one_nontrivial hχ₀, zero_mul, add_zero] + · rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] + have H := MulChar.exists_apply_sub_one_mul_apply_sub_one (by omega) hχ hψ hμ + have Hcs x := (H x).choose_spec + refine ⟨-q * z₁ + ∑ x ∈ (univ \ {0, 1} : Finset F), (H x).choose, ?_, ?_⟩ + · refine Subalgebra.add_mem _ (Subalgebra.mul_mem _ (Subalgebra.neg_mem _ ?_) hz₁) ?_ + · exact Subalgebra.natCast_mem .. + · exact Subalgebra.sum_mem _ fun x _ ↦ (Hcs x).1 + · conv => enter [1, 2, 2, x]; rw [(Hcs x).2] + rw [← Finset.sum_mul, Nat.cast_add, Nat.cast_mul, Hz₁] + ring + +end image From fbc97bc51ebfd507dc49a8cc08fe1a79e9284362 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 1 Oct 2024 14:00:38 +0000 Subject: [PATCH 084/174] chore: split longFile test to its own file (#17311) The longFile test reports line numbers and changes to the tests in the same file trigger changes in these error messages that are really unrelated. Moving the tests for the longFile to a separate file prevents this behaviour. --- test/Lint.lean | 67 --------------------------------------------- test/LongFile.lean | 68 ++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 68 insertions(+), 67 deletions(-) create mode 100644 test/LongFile.lean diff --git a/test/Lint.lean b/test/Lint.lean index c4f49cc99687e..25bf9c9147d13 100644 --- a/test/Lint.lean +++ b/test/Lint.lean @@ -294,70 +294,3 @@ note: this linter can be disabled with `set_option linter.style.longLine false` #guard_msgs in set_option linter.style.longLine true in #check " \" " - -/- -# Testing the `longFile` linter - -Things to note: -* `set_option linter.style.longFile 0` disables the linter, allowing us to set a value smaller than - `1500` without triggering the warning for setting a small value for the option; -* `guard_msgs ... in #exit` and `set_option ... in #exit` allow processing of the file *beyond* - `#exit`, since they wrap `#exit` inside an anonymous section, - making Lean active again *after* that anonymous section. - --/ - -section longFile - -/-- -warning: The default value of the `longFile` linter is 1500. -The current value of 1500 does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1500`. --/ -#guard_msgs in --- Do not allow setting a "small" `longFile` linter option -set_option linter.style.longFile 1500 - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: The default value of the `longFile` linter is 1500. -This file is 331 lines long which does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1600`. --/ -#guard_msgs in --- Do not allow unnecessarily increasing the `longFile` linter option -set_option linter.style.longFile 1600 in -#exit - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: This file is 346 lines long, but the limit is 10. - -You can extend the allowed length of the file using `set_option linter.style.longFile 1500`. -You can completely disable this linter by setting the length limit to `0`. --/ -#guard_msgs in --- First, we silence the linter, so that we can set a default value smaller than 1500. -set_option linter.style.longFile 0 in --- Next, we test that the `longFile` linter warns when a file exceeds the allowed value. -set_option linter.style.longFile 10 in -#exit - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: The default value of the `longFile` linter is 1500. -This file is 361 lines long which does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1700`. --/ -#guard_msgs in --- First, we silence the linter, so that we can set a default value smaller than 1500. -set_option linter.style.longFile 0 in --- If we set the allowed bound for the `longFile` linter that is too large, --- the linter tells us to use a smaller bound. -set_option linter.style.longFile 1700 in -#exit - -end longFile diff --git a/test/LongFile.lean b/test/LongFile.lean new file mode 100644 index 0000000000000..faec16c404806 --- /dev/null +++ b/test/LongFile.lean @@ -0,0 +1,68 @@ +import Mathlib.Tactic.Linter.Lint + +/- +# Testing the `longFile` linter + +Things to note: +* `set_option linter.style.longFile 0` disables the linter, allowing us to set a value smaller than + `1500` without triggering the warning for setting a small value for the option; +* `guard_msgs ... in #exit` and `set_option ... in #exit` allow processing of the file *beyond* + `#exit`, since they wrap `#exit` inside an anonymous section, + making Lean active again *after* that anonymous section. + +-/ + +section longFile + +/-- +warning: The default value of the `longFile` linter is 1500. +The current value of 1500 does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1500`. +-/ +#guard_msgs in +-- Do not allow setting a "small" `longFile` linter option +set_option linter.style.longFile 1500 + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: The default value of the `longFile` linter is 1500. +This file is 36 lines long which does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1600`. +-/ +#guard_msgs in +-- Do not allow unnecessarily increasing the `longFile` linter option +set_option linter.style.longFile 1600 in +#exit + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: This file is 51 lines long, but the limit is 10. + +You can extend the allowed length of the file using `set_option linter.style.longFile 1500`. +You can completely disable this linter by setting the length limit to `0`. +-/ +#guard_msgs in +-- First, we silence the linter, so that we can set a default value smaller than 1500. +set_option linter.style.longFile 0 in +-- Next, we test that the `longFile` linter warns when a file exceeds the allowed value. +set_option linter.style.longFile 10 in +#exit + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: The default value of the `longFile` linter is 1500. +This file is 66 lines long which does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1700`. +-/ +#guard_msgs in +-- First, we silence the linter, so that we can set a default value smaller than 1500. +set_option linter.style.longFile 0 in +-- If we set the allowed bound for the `longFile` linter that is too large, +-- the linter tells us to use a smaller bound. +set_option linter.style.longFile 1700 in +#exit + +end longFile From f73be5b0ca97f80cde7c032e61b68b20a5835794 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 1 Oct 2024 14:00:40 +0000 Subject: [PATCH 085/174] feat: add option for the default length of a file (#17312) As discussed on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Retrieving.20.60weak.60.20options/near/468301685). A further benefit of having the option is that the tests can be performed more easily, since you do not have to reach a file length of more than 1500 to test behaviour beyond the `defValue`. --- Mathlib/Tactic/Linter/Lint.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index e1345febe9b36..6d00f05354dde 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -320,6 +320,12 @@ register_option linter.style.longFile : Nat := { descr := "enable the longFile linter" } +/-- The number of lines that the `longFile` linter considers the default. -/ +register_option linter.style.longFileDefValue : Nat := { + defValue := 1500 + descr := "a soft upper bound on the number of lines of each file" +} + namespace Style.longFile @[inherit_doc Mathlib.Linter.linter.style.longFile] @@ -327,7 +333,7 @@ def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do let linterBound := linter.style.longFile.get (← getOptions) if linterBound == 0 then return - let defValue := 1500 + let defValue := linter.style.longFileDefValue.get (← getOptions) let smallOption := match stx with | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue | _ => false From a9d07e8c4a6f1aea4af92e241deabf82d8dc1439 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Tue, 1 Oct 2024 14:56:24 +0000 Subject: [PATCH 086/174] feat(RingTheory/IsTensorProduct): cancel pushout square on the left (#17028) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a commutative diagram of rings ``` R → S → T ↓ ↓ ↓ R' → S' → T' ``` where the left-hand square is a pushout and the big rectangle is a pushout, we show that then also the right-hand square is a pushout. --- Mathlib/Algebra/Module/LinearMap/Defs.lean | 21 ++++++-- Mathlib/RingTheory/IsTensorProduct.lean | 58 ++++++++++++++++++++++ 2 files changed, 75 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index e442944e3cb7b..e2c0d89223846 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -933,9 +933,8 @@ end Actions section RestrictScalarsAsLinearMap -variable {R S M N : Type*} [Semiring R] [Semiring S] [AddCommGroup M] [AddCommGroup N] [Module R M] - [Module R N] [Module S M] [Module S N] - [LinearMap.CompatibleSMul M N R S] +variable {R S M N P : Type*} [Semiring R] [Semiring S] [AddCommMonoid M] [AddCommMonoid N] + [Module R M] [Module R N] [Module S M] [Module S N] [CompatibleSMul M N R S] variable (R S M N) in @[simp] @@ -948,7 +947,9 @@ theorem restrictScalars_add (f g : M →ₗ[S] N) : rfl @[simp] -theorem restrictScalars_neg (f : M →ₗ[S] N) : (-f).restrictScalars R = -f.restrictScalars R := +theorem restrictScalars_neg {M N : Type*} [AddCommGroup M] [AddCommGroup N] + [Module R M] [Module R N] [Module S M] [Module S N] [CompatibleSMul M N R S] + (f : M →ₗ[S] N) : (-f).restrictScalars R = -f.restrictScalars R := rfl variable {R₁ : Type*} [Semiring R₁] [Module R₁ N] [SMulCommClass S R₁ N] [SMulCommClass R R₁ N] @@ -958,6 +959,18 @@ theorem restrictScalars_smul (c : R₁) (f : M →ₗ[S] N) : (c • f).restrictScalars R = c • f.restrictScalars R := rfl +@[simp] +lemma restrictScalars_comp [AddCommMonoid P] [Module S P] [Module R P] + [CompatibleSMul N P R S] [CompatibleSMul M P R S] (f : N →ₗ[S] P) (g : M →ₗ[S] N) : + (f ∘ₗ g).restrictScalars R = f.restrictScalars R ∘ₗ g.restrictScalars R := by + rfl + +@[simp] +lemma restrictScalars_trans {T : Type*} [CommSemiring T] [Module T M] [Module T N] + [CompatibleSMul M N S T] [CompatibleSMul M N R T] (f : M →ₗ[T] N) : + (f.restrictScalars S).restrictScalars R = f.restrictScalars R := + rfl + variable (S M N R R₁) /-- `LinearMap.restrictScalars` as a `LinearMap`. -/ diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index d146835dbbfcd..b02662937d43f 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -318,11 +318,40 @@ theorem IsBaseChange.comp {f : M →ₗ[R] N} (hf : IsBaseChange S f) {g : N → ext rfl +/-- If `N` is the base change of `M` to `S` and `O` the base change of `M` to `T`, then +`O` is the base change of `N` to `T`. -/ +lemma IsBaseChange.of_comp {f : M →ₗ[R] N} (hf : IsBaseChange S f) {h : N →ₗ[S] O} + (hc : IsBaseChange T ((h : N →ₗ[R] O) ∘ₗ f)) : + IsBaseChange T h := by + apply IsBaseChange.of_lift_unique + intro Q _ _ _ _ r + letI : Module R Q := inferInstanceAs (Module R (RestrictScalars R S Q)) + haveI : IsScalarTower R S Q := IsScalarTower.of_algebraMap_smul fun r ↦ congrFun rfl + haveI : IsScalarTower R T Q := IsScalarTower.of_algebraMap_smul fun r x ↦ by + simp [IsScalarTower.algebraMap_apply R S T] + let r' : M →ₗ[R] Q := r ∘ₗ f + let q : O →ₗ[T] Q := hc.lift r' + refine ⟨q, ?_, ?_⟩ + · apply hf.algHom_ext' + simp [LinearMap.comp_assoc, hc.lift_comp] + · intro q' hq' + apply hc.algHom_ext' + apply_fun LinearMap.restrictScalars R at hq' + rw [← LinearMap.comp_assoc] + rw [show q'.restrictScalars R ∘ₗ h.restrictScalars R = _ from hq', hc.lift_comp] + +/-- If `N` is the base change `M` to `S`, then `O` is the base change of `M` to `T` if and +only if `O` is the base change of `N` to `T`. -/ +lemma IsBaseChange.comp_iff {f : M →ₗ[R] N} (hf : IsBaseChange S f) {h : N →ₗ[S] O} : + IsBaseChange T ((h : N →ₗ[R] O) ∘ₗ f) ↔ IsBaseChange T h := + ⟨fun hc ↦ IsBaseChange.of_comp hf hc, fun hh ↦ IsBaseChange.comp hf hh⟩ + variable {R' S' : Type*} [CommSemiring R'] [CommSemiring S'] variable [Algebra R R'] [Algebra S S'] [Algebra R' S'] [Algebra R S'] variable [IsScalarTower R R' S'] [IsScalarTower R S S'] open IsScalarTower (toAlgHom) +open IsScalarTower (algebraMap_apply) variable (R S R' S') @@ -471,4 +500,33 @@ theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type · intro s₁ s₂ e₁ e₂ rw [map_add, map_add, e₁, e₂] +/-- +Let the following be a commutative diagram of rings +``` + R → S → T + ↓ ↓ ↓ + R' → S' → T' +``` +where the left-hand square is a pushout. Then the following are equivalent: +- the big rectangle is a pushout. +- the right-hand square is a pushout. + +Note that this is essentially the isomorphism `T ⊗[S] (S ⊗[R] R') ≃ₐ[T] T ⊗[R] R'`. +-/ +lemma Algebra.IsPushout.comp_iff {T' : Type*} [CommRing T'] [Algebra R T'] + [Algebra S' T'] [Algebra S T'] [Algebra T T'] [Algebra R' T'] + [IsScalarTower R T T'] [IsScalarTower S T T'] [IsScalarTower S S' T'] + [IsScalarTower R R' T'] [IsScalarTower R S' T'] [IsScalarTower R' S' T'] + [Algebra.IsPushout R S R' S'] : + Algebra.IsPushout R T R' T' ↔ Algebra.IsPushout S T S' T' := by + let f : R' →ₗ[R] S' := (IsScalarTower.toAlgHom R R' S').toLinearMap + haveI : IsScalarTower R S T' := IsScalarTower.of_algebraMap_eq <| fun x ↦ by + rw [algebraMap_apply R S' T', algebraMap_apply R S S', ← algebraMap_apply S S' T'] + have heq : (toAlgHom S S' T').toLinearMap.restrictScalars R ∘ₗ f = + (toAlgHom R R' T').toLinearMap := by + ext x + simp [f, ← IsScalarTower.algebraMap_apply] + rw [isPushout_iff, isPushout_iff, ← heq, IsBaseChange.comp_iff] + exact Algebra.IsPushout.out + end IsBaseChange From 2ea60e1661f202eebe4c3fd1abb766198ac4d91b Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 1 Oct 2024 16:18:07 +0000 Subject: [PATCH 087/174] chore: update Mathlib dependencies 2024-10-01 (#17328) 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 6c6e4f48601c1..9161d8bacca4b 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "28fa80508edc97d96ed6342c9a771a67189e0baa", + "rev": "ff420521a0c098891f4f44ecda9dd7ff57b50bad", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4261c7c9290da3471548a327a4fa9387e9fdd684", + "rev": "c9e106b0541a3b1c4bf82017aca78a50cc3e5ca2", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From c9ef6fe03a937ddd560ce766c90036311ae12429 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Tue, 1 Oct 2024 17:04:19 +0000 Subject: [PATCH 088/174] feat(RingTheory/StandardSmooth): composition and localization (#16827) Constructs submersive presentations for composition and localization away from one element. This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024. --- Mathlib/Algebra/MvPolynomial/Basic.lean | 16 ++ Mathlib/Algebra/MvPolynomial/PDeriv.lean | 21 ++ Mathlib/RingTheory/Presentation.lean | 15 +- Mathlib/RingTheory/Smooth/StandardSmooth.lean | 236 +++++++++++++++++- 4 files changed, 277 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index da687de7c7942..722c1bdf7b69a 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -1357,6 +1357,11 @@ theorem comp_aeval {B : Type*} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[ ext i simp +lemma comp_aeval_apply {B : Type*} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[R] B) + (p : MvPolynomial σ R) : + φ (aeval f p) = aeval (fun i ↦ φ (f i)) p := by + rw [← comp_aeval, AlgHom.coe_comp, comp_apply] + @[simp] theorem map_aeval {B : Type*} [CommSemiring B] (g : σ → S₁) (φ : S₁ →+* B) (p : MvPolynomial σ R) : φ (aeval g p) = eval₂Hom (φ.comp (algebraMap R S₁)) (fun i => φ (g i)) p := by @@ -1537,6 +1542,17 @@ theorem eval_mem {p : MvPolynomial σ S} {s : subS} (hs : ∀ i ∈ p.support, p end EvalMem +variable {S T : Type*} [CommSemiring S] [Algebra R S] [CommSemiring T] [Algebra R T] [Algebra S T] + [IsScalarTower R S T] + +lemma aeval_sum_elim {σ τ : Type*} (p : MvPolynomial (σ ⊕ τ) R) (f : τ → S) (g : σ → T) : + (aeval (Sum.elim g (algebraMap S T ∘ f))) p = + (aeval g) ((aeval (Sum.elim X (C ∘ f))) p) := by + induction' p using MvPolynomial.induction_on with r p q hp hq p i h + · simp [← IsScalarTower.algebraMap_apply] + · simp [hp, hq] + · cases i <;> simp [h] + end CommSemiring end MvPolynomial diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index a4755788c2c2a..b7aa67a2aad48 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -115,6 +115,27 @@ theorem pderiv_map {S} [CommSemiring S] {φ : R →+* S} {f : MvPolynomial σ R} · simp [eq] · simp [eq, h] +lemma pderiv_rename {τ : Type*} [DecidableEq τ] [DecidableEq σ] {f : σ → τ} + (hf : Function.Injective f) (x : σ) (p : MvPolynomial σ R) : + pderiv (f x) (rename f p) = rename f (pderiv x p) := by + induction' p using MvPolynomial.induction_on with a p q hp hq p a h + · simp + · simp [hp, hq] + · simp only [map_mul, MvPolynomial.rename_X, Derivation.leibniz, MvPolynomial.pderiv_X, + Pi.single_apply, hf.eq_iff, smul_eq_mul, mul_ite, mul_one, mul_zero, h, map_add, add_left_inj] + split_ifs <;> simp + +lemma aeval_sum_elim_pderiv_inl {S τ : Type*} [CommRing S] [Algebra R S] + (p : MvPolynomial (σ ⊕ τ) R) (f : τ → S) (j : σ) : + aeval (Sum.elim X (C ∘ f)) ((pderiv (Sum.inl j)) p) = + (pderiv j) ((aeval (Sum.elim X (C ∘ f))) p) := by + classical + induction' p using MvPolynomial.induction_on with a p q hp hq p q h + · simp + · simp [hp, hq] + · simp only [Derivation.leibniz, pderiv_X, smul_eq_mul, map_add, map_mul, aeval_X, h] + cases q <;> simp [Pi.single_apply] + end PDeriv end MvPolynomial diff --git a/Mathlib/RingTheory/Presentation.lean b/Mathlib/RingTheory/Presentation.lean index 1fbea429322dc..2cdda618659b5 100644 --- a/Mathlib/RingTheory/Presentation.lean +++ b/Mathlib/RingTheory/Presentation.lean @@ -179,6 +179,7 @@ private lemma span_range_relation_eq_ker_localizationAway : show Ideal.span {C r * X () - 1} = Ideal.comap _ (RingHom.ker (mvPolynomialQuotientEquiv S r)) simp [RingHom.ker_equiv, ← RingHom.ker_eq_comap_bot] +variable (S) in /-- If `S` is the localization of `R` away from `r`, we can construct a natural presentation of `S` as `R`-algebra with a single generator `X` and the relation `r * X - 1 = 0`. -/ @[simps relation, simps (config := .lemmasOnly) rels] @@ -190,12 +191,15 @@ noncomputable def localizationAway : Presentation R S where simp only [Generators.localizationAway_vars, Set.range_const] apply span_range_relation_eq_ker_localizationAway r -instance localizationAway_isFinite : (localizationAway r (S := S)).IsFinite where +instance localizationAway_isFinite : (localizationAway S r).IsFinite where finite_vars := inferInstanceAs <| Finite Unit finite_rels := inferInstanceAs <| Finite Unit +instance : Fintype (localizationAway S r).rels := + inferInstanceAs (Fintype Unit) + @[simp] -lemma localizationAway_dimension_zero : (localizationAway r (S := S)).dimension = 0 := by +lemma localizationAway_dimension_zero : (localizationAway S r).dimension = 0 := by simp [Presentation.dimension, localizationAway, Generators.localizationAway_vars] end Localization @@ -408,7 +412,12 @@ noncomputable def comp : Presentation R T where (fun rp ↦ MvPolynomial.rename Sum.inr <| P.relation rp) span_range_relation_eq_ker := Q.span_range_relation_eq_ker_comp P -lemma comp_relation_map (r : Q.rels) : +@[simp] +lemma comp_relation_inr (r : P.rels) : + (Q.comp P).relation (Sum.inr r) = rename Sum.inr (P.relation r) := + rfl + +lemma comp_aeval_relation_inl (r : Q.rels) : aeval (Sum.elim X (MvPolynomial.C ∘ P.val)) ((Q.comp P).relation (Sum.inl r)) = Q.relation r := by show (Q.aux P) _ = _ diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index d8064eb6a7e8c..f9ed2c1090626 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -62,10 +62,6 @@ Finally, for ring homomorphisms we define: ## TODO -- Show that the canonical presentation for localization away from an element is standard smooth - of relative dimension 0. -- Show that the composition of submersive presentations of relative dimensions `n` and `m` is - submersive of relative dimension `n + m`. - Show that the module of Kaehler differentials of a standard smooth `R`-algebra `S` of relative dimension `n` is `S`-free of rank `n`. In particular this shows that the relative dimension is independent of the choice of the standard smooth presentation. @@ -87,9 +83,9 @@ in June 2024. universe t t' w w' u v -open TensorProduct Classical +open TensorProduct MvPolynomial Classical -variable (n : ℕ) +variable (n m : ℕ) namespace Algebra @@ -161,7 +157,7 @@ lemma jacobian_eq_jacobiMatrix_det : P.jacobian = algebraMap P.Ring S P.jacobiMa lemma jacobiMatrix_apply (i j : P.rels) : P.jacobiMatrix i j = MvPolynomial.pderiv (P.map i) (P.relation j) := by - simp [jacobiMatrix, LinearMap.toMatrix, differential] + simp [jacobiMatrix, LinearMap.toMatrix, differential, basis] end Matrix @@ -182,6 +178,7 @@ instance (h : Function.Bijective (algebraMap R S)) : Fintype (ofBijectiveAlgebra instance (h : Function.Bijective (algebraMap R S)) : Fintype (ofBijectiveAlgebraMap h).rels := inferInstanceAs (Fintype PEmpty) +@[simp] lemma ofBijectiveAlgebraMap_jacobian (h : Function.Bijective (algebraMap R S)) : (ofBijectiveAlgebraMap h).jacobian = 1 := by have : (algebraMap (ofBijectiveAlgebraMap h).Ring S).mapMatrix @@ -190,7 +187,167 @@ lemma ofBijectiveAlgebraMap_jacobian (h : Function.Bijective (algebraMap R S)) : contradiction rw [jacobian_eq_jacobiMatrix_det, RingHom.map_det, this, Matrix.det_one] -open MvPolynomial +section Localization + +variable (r : R) [IsLocalization.Away r S] + +variable (S) in +/-- If `S` is the localization of `R` at `r`, this is the canonical submersive presentation +of `S` as `R`-algebra. -/ +@[simps map] +noncomputable def localizationAway : PreSubmersivePresentation R S where + __ := Presentation.localizationAway S r + map _ := () + map_inj _ _ h := h + relations_finite := inferInstanceAs <| Finite Unit + +instance : Fintype (localizationAway S r).rels := + inferInstanceAs (Fintype Unit) + +instance : DecidableEq (localizationAway S r).rels := + inferInstanceAs (DecidableEq Unit) + +@[simp] +lemma localizationAway_jacobiMatrix : + (localizationAway S r).jacobiMatrix = Matrix.diagonal (fun () ↦ MvPolynomial.C r) := by + have h : (pderiv ()) (C r * X () - 1) = C r := by simp + ext (i : Unit) (j : Unit) : 1 + rwa [jacobiMatrix_apply] + +@[simp] +lemma localizationAway_jacobian : (localizationAway S r).jacobian = algebraMap R S r := by + rw [jacobian_eq_jacobiMatrix_det, localizationAway_jacobiMatrix] + simp [show Fintype.card (localizationAway r (S := S)).rels = 1 from rfl] + +end Localization + +section Composition + +variable {T} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable (Q : PreSubmersivePresentation S T) (P : PreSubmersivePresentation R S) + +/-- Given an `R`-algebra `S` and an `S`-algebra `T` with pre-submersive presentations, +this is the canonical pre-submersive presentation of `T` as an `R`-algebra. -/ +@[simps map] +noncomputable def comp : PreSubmersivePresentation R T where + __ := Q.toPresentation.comp P.toPresentation + map := Sum.elim (fun rq ↦ Sum.inl <| Q.map rq) (fun rp ↦ Sum.inr <| P.map rp) + map_inj := Function.Injective.sum_elim ((Sum.inl_injective).comp (Q.map_inj)) + ((Sum.inr_injective).comp (P.map_inj)) <| by simp + relations_finite := inferInstanceAs <| Finite (Q.rels ⊕ P.rels) + +/-- The dimension of the composition of two finite submersive presentations is +the sum of the dimensions. -/ +lemma dimension_comp_eq_dimension_add_dimension [Q.IsFinite] [P.IsFinite] : + (Q.comp P).dimension = Q.dimension + P.dimension := by + simp only [Presentation.dimension] + erw [Presentation.comp_rels, Generators.comp_vars] + have : Nat.card P.rels ≤ Nat.card P.vars := + card_relations_le_card_vars_of_isFinite P + have : Nat.card Q.rels ≤ Nat.card Q.vars := + card_relations_le_card_vars_of_isFinite Q + simp only [Nat.card_sum] + omega + +section + +/-! +### Jacobian of composition + +Let `S` be an `R`-algebra and `T` be an `S`-algebra with presentations `P` and `Q` respectively. +In this section we compute the jacobian of the composition of `Q` and `P` to be +the product of the jacobians. For this we use a block decomposition of the jacobi matrix and show +that the upper-right block vanishes, the upper-left block has determinant jacobian of `Q` and +the lower-right block has determinant jacobian of `P`. + +-/ + +variable [Fintype (Q.comp P).rels] + +private lemma jacobiMatrix_comp_inl_inr (i : Q.rels) (j : P.rels) : + (Q.comp P).jacobiMatrix (Sum.inl i) (Sum.inr j) = 0 := by + rw [jacobiMatrix_apply] + refine MvPolynomial.pderiv_eq_zero_of_not_mem_vars (fun hmem ↦ ?_) + apply MvPolynomial.vars_rename at hmem + simp at hmem + +private lemma jacobiMatrix_comp_₁₂ : (Q.comp P).jacobiMatrix.toBlocks₁₂ = 0 := by + ext i j : 1 + simp [Matrix.toBlocks₁₂, jacobiMatrix_comp_inl_inr] + +section Q + +variable [Fintype Q.rels] + +private lemma jacobiMatrix_comp_inl_inl (i j : Q.rels) : + aeval (Sum.elim X (MvPolynomial.C ∘ P.val)) + ((Q.comp P).jacobiMatrix (Sum.inl j) (Sum.inl i)) = Q.jacobiMatrix j i := by + rw [jacobiMatrix_apply, jacobiMatrix_apply, comp_map, Sum.elim_inl, + ← Q.comp_aeval_relation_inl P.toPresentation] + apply aeval_sum_elim_pderiv_inl + +private lemma jacobiMatrix_comp_₁₁_det : + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₁₁.det = Q.jacobian := by + rw [jacobian_eq_jacobiMatrix_det, AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det] + congr + ext i j : 1 + simp only [Matrix.map_apply, RingHom.mapMatrix_apply, ← Q.jacobiMatrix_comp_inl_inl P] + apply aeval_sum_elim + +end Q + +section P + +variable [Fintype P.rels] + +private lemma jacobiMatrix_comp_inr_inr (i j : P.rels) : + (Q.comp P).jacobiMatrix (Sum.inr i) (Sum.inr j) = + MvPolynomial.rename Sum.inr (P.jacobiMatrix i j) := by + rw [jacobiMatrix_apply, jacobiMatrix_apply] + simp only [comp_map, Sum.elim_inr] + apply pderiv_rename Sum.inr_injective + +private lemma jacobiMatrix_comp_₂₂_det : + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₂₂.det = algebraMap S T P.jacobian := by + rw [jacobian_eq_jacobiMatrix_det] + rw [AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det, RingHom.map_det] + congr + ext i j : 1 + simp only [Matrix.toBlocks₂₂, AlgHom.mapMatrix_apply, Matrix.map_apply, Matrix.of_apply, + RingHom.mapMatrix_apply, Generators.algebraMap_apply, map_aeval, coe_eval₂Hom] + rw [jacobiMatrix_comp_inr_inr, ← IsScalarTower.algebraMap_eq] + simp only [aeval, AlgHom.coe_mk, coe_eval₂Hom] + generalize P.jacobiMatrix i j = p + induction' p using MvPolynomial.induction_on with a p q hp hq p i hp + · simp only [algHom_C, algebraMap_eq, eval₂_C] + erw [MvPolynomial.eval₂_C] + · simp [hp, hq] + · simp only [map_mul, rename_X, eval₂_mul, hp, eval₂_X] + erw [Generators.comp_val] + simp + +end P + +end + +/-- The jacobian of the composition of presentations is the product of the jacobians. -/ +@[simp] +lemma comp_jacobian_eq_jacobian_smul_jacobian : (Q.comp P).jacobian = P.jacobian • Q.jacobian := by + cases nonempty_fintype Q.rels + cases nonempty_fintype P.rels + letI : Fintype (Q.comp P).rels := inferInstanceAs <| Fintype (Q.rels ⊕ P.rels) + rw [jacobian_eq_jacobiMatrix_det, ← Matrix.fromBlocks_toBlocks ((Q.comp P).jacobiMatrix), + jacobiMatrix_comp_₁₂] + convert_to + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₁₁.det * + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₂₂.det = P.jacobian • Q.jacobian + · simp only [Generators.algebraMap_apply, ← map_mul] + congr + convert Matrix.det_fromBlocks_zero₁₂ (Q.comp P).jacobiMatrix.toBlocks₁₁ + (Q.comp P).jacobiMatrix.toBlocks₂₁ (Q.comp P).jacobiMatrix.toBlocks₂₂ + · rw [jacobiMatrix_comp_₁₁_det, jacobiMatrix_comp_₂₂_det, mul_comm, Algebra.smul_def] + +end Composition section BaseChange @@ -204,6 +361,7 @@ noncomputable def baseChange : PreSubmersivePresentation T (T ⊗[R] S) where map_inj := P.map_inj relations_finite := P.relations_finite +@[simp] lemma baseChange_jacobian : (P.baseChange T).jacobian = 1 ⊗ₜ P.jacobian := by classical cases nonempty_fintype P.rels @@ -258,6 +416,37 @@ noncomputable def ofBijectiveAlgebraMap (h : Function.Bijective (algebraMap R S) noncomputable def id : SubmersivePresentation.{t, w} R R := ofBijectiveAlgebraMap Function.bijective_id +section Composition + +variable {R S T} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable (Q : SubmersivePresentation S T) (P : SubmersivePresentation R S) + +/-- Given an `R`-algebra `S` and an `S`-algebra `T` with submersive presentations, +this is the canonical submersive presentation of `T` as an `R`-algebra. -/ +noncomputable def comp : SubmersivePresentation R T where + __ := Q.toPreSubmersivePresentation.comp P.toPreSubmersivePresentation + jacobian_isUnit := by + rw [comp_jacobian_eq_jacobian_smul_jacobian, Algebra.smul_def, IsUnit.mul_iff] + exact ⟨RingHom.isUnit_map _ <| P.jacobian_isUnit, Q.jacobian_isUnit⟩ + isFinite := Presentation.comp_isFinite Q.toPresentation P.toPresentation + +end Composition + +section Localization + +variable {R} (r : R) [IsLocalization.Away r S] + +/-- If `S` is the localization of `R` at `r`, this is the canonical submersive presentation +of `S` as `R`-algebra. -/ +noncomputable def localizationAway : SubmersivePresentation R S where + __ := PreSubmersivePresentation.localizationAway S r + jacobian_isUnit := by + rw [localizationAway_jacobian] + apply IsLocalization.map_units' (⟨r, 1, by simp⟩ : Submonoid.powers r) + isFinite := Presentation.localizationAway_isFinite r + +end Localization + section BaseChange variable (T) [CommRing T] [Algebra R T] (P : SubmersivePresentation R S) @@ -316,6 +505,37 @@ instance IsStandardSmoothOfRelativeDimension.id : IsStandardSmoothOfRelativeDimension.{t, w} 0 R R := IsStandardSmoothOfRelativeDimension.of_algebraMap_bijective Function.bijective_id +section Composition + +variable (R S T) [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + +lemma IsStandardSmooth.trans [IsStandardSmooth.{t, w} R S] [IsStandardSmooth.{t', w'} S T] : + IsStandardSmooth.{max t t', max w w'} R T where + out := by + obtain ⟨⟨P⟩⟩ := ‹IsStandardSmooth R S› + obtain ⟨⟨Q⟩⟩ := ‹IsStandardSmooth S T› + exact ⟨Q.comp P⟩ + +lemma IsStandardSmoothOfRelativeDimension.trans [IsStandardSmoothOfRelativeDimension.{t, w} n R S] + [IsStandardSmoothOfRelativeDimension.{t', w'} m S T] : + IsStandardSmoothOfRelativeDimension.{max t t', max w w'} (m + n) R T where + out := by + obtain ⟨P, hP⟩ := ‹IsStandardSmoothOfRelativeDimension n R S› + obtain ⟨Q, hQ⟩ := ‹IsStandardSmoothOfRelativeDimension m S T› + refine ⟨Q.comp P, hP ▸ hQ ▸ ?_⟩ + apply PreSubmersivePresentation.dimension_comp_eq_dimension_add_dimension + +end Composition + +lemma IsStandardSmooth.localization_away (r : R) [IsLocalization.Away r S] : + IsStandardSmooth.{0, 0} R S where + out := ⟨SubmersivePresentation.localizationAway S r⟩ + +lemma IsStandardSmoothOfRelativeDimension.localization_away (r : R) [IsLocalization.Away r S] : + IsStandardSmoothOfRelativeDimension.{0, 0} 0 R S where + out := ⟨SubmersivePresentation.localizationAway S r, + Presentation.localizationAway_dimension_zero r⟩ + section BaseChange variable (T) [CommRing T] [Algebra R T] From 0d57a343f3023f6ef0c0a81f857534fcb34ebedc Mon Sep 17 00:00:00 2001 From: Raph-DG Date: Tue, 1 Oct 2024 17:04:20 +0000 Subject: [PATCH 089/174] feat(Topology/KrullDimension): add invariance under homeomorphism theorem and auxiliary lemmas (#17175) Added some lemmas about the topological Krull dimension --- Mathlib/Data/Set/Image.lean | 3 +++ Mathlib/Topology/KrullDimension.lean | 39 ++++++++++++++++++++++++++++ 2 files changed, 42 insertions(+) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 9d317af2926e5..9b390e67bad19 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -1090,6 +1090,9 @@ theorem Injective.image_injective (hf : Injective f) : Injective (image f) := by intro s t h rw [← preimage_image_eq s hf, ← preimage_image_eq t hf, h] +lemma Injective.image_strictMono (inj : Function.Injective f) : StrictMono (image f) := + monotone_image.strictMono_of_injective inj.image_injective + theorem Surjective.preimage_subset_preimage_iff {s t : Set β} (hf : Surjective f) : f ⁻¹' s ⊆ f ⁻¹' t ↔ s ⊆ t := by apply Set.preimage_subset_preimage_iff diff --git a/Mathlib/Topology/KrullDimension.lean b/Mathlib/Topology/KrullDimension.lean index ea4e77fb0194c..b5eee2e982ae4 100644 --- a/Mathlib/Topology/KrullDimension.lean +++ b/Mathlib/Topology/KrullDimension.lean @@ -22,3 +22,42 @@ closed irreducible sets. -/ noncomputable def topologicalKrullDim (T : Type*) [TopologicalSpace T] : WithBot ℕ∞ := krullDim (IrreducibleCloseds T) + +variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + +/-- +Map induced on irreducible closed subsets by a closed continuous map `f`. +This is just a wrapper around the image of `f` together with proofs that it +preserves irreducibility (by continuity) and closedness (since `f` is closed). +-/ +def IrreducibleCloseds.map {f : X → Y} (hf1 : Continuous f) (hf2 : IsClosedMap f) + (c : IrreducibleCloseds X) : + IrreducibleCloseds Y where + carrier := f '' c + is_irreducible' := c.is_irreducible'.image f hf1.continuousOn + is_closed' := hf2 c c.is_closed' + +/-- +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) : + StrictMono (IrreducibleCloseds.map hf.continuous hf.isClosedMap) := + fun ⦃_ _⦄ UltV ↦ hf.inj.image_strictMono UltV + +/-- +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) : + topologicalKrullDim X ≤ topologicalKrullDim Y := + krullDim_le_of_strictMono _ (IrreducibleCloseds.map_strictMono hf) + +/-- 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 + have bwd : topologicalKrullDim Y ≤ topologicalKrullDim X := + ClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm + (h.homeomorph f).symm.closedEmbedding + le_antisymm fwd bwd From 1fe4aec4735318982df91ad8ff0581406c6b625e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 17:04:21 +0000 Subject: [PATCH 090/174] =?UTF-8?q?feat:=20`=E2=84=82`=20and=20`Finset.exp?= =?UTF-8?q?ect`=20(#17274)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Make more arguments to `NNReal.coe_sum` explicit to match `NNReal.coe_expect`. From LeanAPAP --- .../Algebra/Order/BigOperators/Expect.lean | 3 +-- Mathlib/Analysis/Normed/Group/Basic.lean | 2 +- .../Analysis/Normed/Group/Constructions.lean | 2 +- Mathlib/Analysis/RCLike/Basic.lean | 25 +++++++++++++++---- Mathlib/Data/Complex/Basic.lean | 9 +++++++ Mathlib/Data/Complex/BigOperators.lean | 17 +++++++++++-- Mathlib/Data/NNReal/Basic.lean | 24 ++++++++++-------- Mathlib/Topology/Instances/Real.lean | 9 ++++--- 8 files changed, 65 insertions(+), 26 deletions(-) diff --git a/Mathlib/Algebra/Order/BigOperators/Expect.lean b/Mathlib/Algebra/Order/BigOperators/Expect.lean index 30064a4bcb627..a449b9ce36435 100644 --- a/Mathlib/Algebra/Order/BigOperators/Expect.lean +++ b/Mathlib/Algebra/Order/BigOperators/Expect.lean @@ -130,8 +130,7 @@ end LinearOrderedAddCommMonoid section LinearOrderedAddCommGroup variable [LinearOrderedAddCommGroup α] [Module ℚ≥0 α] [PosSMulMono ℚ≥0 α] --- TODO: Norm version -lemma abs_expect_le_expect_abs (s : Finset ι) (f : ι → α) : |𝔼 i ∈ s, f i| ≤ 𝔼 i ∈ s, |f i| := +lemma abs_expect_le (s : Finset ι) (f : ι → α) : |𝔼 i ∈ s, f i| ≤ 𝔼 i ∈ s, |f i| := le_expect_of_subadditive abs_zero abs_add (fun _ ↦ abs_nnqsmul _) end LinearOrderedAddCommGroup diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index bde983756cf20..1eba1978859f4 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1129,7 +1129,7 @@ theorem nnnorm_prod_le (s : Finset ι) (f : ι → E) : ‖∏ a ∈ s, f a‖ @[to_additive] theorem nnnorm_prod_le_of_le (s : Finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ‖f b‖₊ ≤ n b) : ‖∏ b ∈ s, f b‖₊ ≤ ∑ b ∈ s, n b := - (norm_prod_le_of_le s h).trans_eq NNReal.coe_sum.symm + (norm_prod_le_of_le s h).trans_eq (NNReal.coe_sum ..).symm namespace Real diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index a43e9e963c9ee..c9ee9cc5d178f 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -362,7 +362,7 @@ lemma Pi.sum_norm_apply_le_norm' : ∑ i, ‖f i‖ ≤ Fintype.card ι • ‖f @[to_additive Pi.sum_nnnorm_apply_le_nnnorm "The $L^1$ norm is less than the $L^\\infty$ norm scaled by the cardinality."] lemma Pi.sum_nnnorm_apply_le_nnnorm' : ∑ i, ‖f i‖₊ ≤ Fintype.card ι • ‖f‖₊ := - NNReal.coe_sum.trans_le <| Pi.sum_norm_apply_le_norm' _ + (NNReal.coe_sum ..).trans_le <| Pi.sum_norm_apply_le_norm' _ end SeminormedGroup diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 869aa9fe0a463..82077c49fead5 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Algebra.Field +import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.Star.Basic import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap @@ -39,7 +40,7 @@ their counterparts in `Mathlib/Analysis/Complex/Basic.lean` (which causes linter A few lemmas requiring heavier imports are in `Mathlib/Data/RCLike/Lemmas.lean`. -/ -open scoped ComplexConjugate +open scoped BigOperators ComplexConjugate section @@ -234,6 +235,10 @@ theorem norm_ofReal (r : ℝ) : ‖(r : K)‖ = |r| := instance (priority := 100) charZero_rclike : CharZero K := (RingHom.charZero_iff (algebraMap ℝ K).injective).1 inferInstance +@[rclike_simps, norm_cast] +lemma ofReal_expect {α : Type*} (s : Finset α) (f : α → ℝ) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : K) := + map_expect (algebraMap ..) .. + /-! ### The imaginary unit, `I` -/ /-- The imaginary unit. -/ @@ -605,13 +610,23 @@ variable (K) in lemma nnnorm_nsmul [NormedAddCommGroup E] [NormedSpace K E] (n : ℕ) (x : E) : ‖n • x‖₊ = n • ‖x‖₊ := by simpa [Nat.cast_smul_eq_nsmul] using nnnorm_smul (n : K) x +section NormedField +variable [NormedField E] [CharZero E] [NormedSpace K E] +include K + variable (K) in -lemma norm_nnqsmul [NormedField E] [CharZero E] [NormedSpace K E] (q : ℚ≥0) (x : E) : - ‖q • x‖ = q • ‖x‖ := by simpa [NNRat.cast_smul_eq_nnqsmul] using norm_smul (q : K) x +lemma norm_nnqsmul (q : ℚ≥0) (x : E) : ‖q • x‖ = q • ‖x‖ := by + simpa [NNRat.cast_smul_eq_nnqsmul] using norm_smul (q : K) x variable (K) in -lemma nnnorm_nnqsmul [NormedField E] [CharZero E] [NormedSpace K E] (q : ℚ≥0) (x : E) : - ‖q • x‖₊ = q • ‖x‖₊ := by simpa [NNRat.cast_smul_eq_nnqsmul] using nnnorm_smul (q : K) x +lemma nnnorm_nnqsmul (q : ℚ≥0) (x : E) : ‖q • x‖₊ = q • ‖x‖₊ := by + simpa [NNRat.cast_smul_eq_nnqsmul] using nnnorm_smul (q : K) x + +@[bound] +lemma norm_expect_le {ι : Type*} {s : Finset ι} {f : ι → E} : ‖𝔼 i ∈ s, f i‖ ≤ 𝔼 i ∈ s, ‖f i‖ := + Finset.le_expect_of_subadditive norm_zero norm_add_le fun _ _ ↦ by rw [norm_nnqsmul K] + +end NormedField theorem mul_self_norm (z : K) : ‖z‖ * ‖z‖ = normSq z := by rw [normSq_eq_def', sq] diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 6483469eea3bc..089de7d56ab00 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -442,6 +442,15 @@ lemma re_ofNat (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℂ).re = @[simp, norm_cast] lemma ratCast_re (q : ℚ) : (q : ℂ).re = q := rfl @[simp, norm_cast] lemma ratCast_im (q : ℚ) : (q : ℂ).im = 0 := rfl +lemma re_nsmul (n : ℕ) (z : ℂ) : (n • z).re = n • z.re := smul_re .. +lemma im_nsmul (n : ℕ) (z : ℂ) : (n • z).im = n • z.im := smul_im .. +lemma re_zsmul (n : ℤ) (z : ℂ) : (n • z).re = n • z.re := smul_re .. +lemma im_zsmul (n : ℤ) (z : ℂ) : (n • z).im = n • z.im := smul_im .. +@[simp] lemma re_nnqsmul (q : ℚ≥0) (z : ℂ) : (q • z).re = q • z.re := smul_re .. +@[simp] lemma im_nnqsmul (q : ℚ≥0) (z : ℂ) : (q • z).im = q • z.im := smul_im .. +@[simp] lemma re_qsmul (q : ℚ) (z : ℂ) : (q • z).re = q • z.re := smul_re .. +@[simp] lemma im_qsmul (q : ℚ) (z : ℂ) : (q • z).im = q • z.im := smul_im .. + @[deprecated (since := "2024-04-17")] alias rat_cast_im := ratCast_im diff --git a/Mathlib/Data/Complex/BigOperators.lean b/Mathlib/Data/Complex/BigOperators.lean index 43c85d7750bfa..cd9cb5f344fa0 100644 --- a/Mathlib/Data/Complex/BigOperators.lean +++ b/Mathlib/Data/Complex/BigOperators.lean @@ -3,14 +3,15 @@ Copyright (c) 2017 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Mario Carneiro -/ -import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.BigOperators.Expect import Mathlib.Data.Complex.Basic /-! # Finite sums and products of complex numbers - -/ +open scoped BigOperators + namespace Complex variable {α : Type*} (s : Finset α) @@ -23,12 +24,24 @@ theorem ofReal_prod (f : α → ℝ) : ((∏ i ∈ s, f i : ℝ) : ℂ) = ∏ i theorem ofReal_sum (f : α → ℝ) : ((∑ i ∈ s, f i : ℝ) : ℂ) = ∑ i ∈ s, (f i : ℂ) := map_sum ofReal _ _ +@[simp, norm_cast] +lemma ofReal_expect (f : α → ℝ) : (𝔼 i ∈ s, f i : ℝ) = 𝔼 i ∈ s, (f i : ℂ) := + map_expect ofReal .. + @[simp] theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re := map_sum reAddGroupHom f s +@[simp] +lemma re_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).re = 𝔼 i ∈ s, (f i).re := + map_expect (LinearMap.mk reAddGroupHom.toAddHom (by simp)) f s + @[simp] theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im := map_sum imAddGroupHom f s +@[simp] +lemma im_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).im = 𝔼 i ∈ s, (f i).im := + map_expect (LinearMap.mk imAddGroupHom.toAddHom (by simp)) f s + end Complex diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index df15ccc645e66..acc5c21138221 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -3,16 +3,13 @@ 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.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Canonical.Basic -import Mathlib.Algebra.Order.Nonneg.Field import Mathlib.Algebra.Order.Nonneg.Floor +import Mathlib.Algebra.Ring.Regular import Mathlib.Data.Real.Pointwise import Mathlib.Order.ConditionallyCompleteLattice.Group -import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.GCongr.CoreAttrs -import Mathlib.Algebra.Ring.Regular /-! # Nonnegative real numbers @@ -55,6 +52,7 @@ This file defines `ℝ≥0` as a localized notation for `NNReal`. assert_not_exists Star open Function +open scoped BigOperators -- to ensure these instances are computable /-- Nonnegative real numbers. -/ @@ -278,22 +276,26 @@ theorem coe_multiset_sum (s : Multiset ℝ≥0) : ((s.sum : ℝ≥0) : ℝ) = (s theorem coe_multiset_prod (s : Multiset ℝ≥0) : ((s.prod : ℝ≥0) : ℝ) = (s.map (↑)).prod := map_multiset_prod toRealHom s +variable {ι : Type*} {s : Finset ι} {f : ι → ℝ} + @[simp, norm_cast] -theorem coe_sum {α} {s : Finset α} {f : α → ℝ≥0} : ↑(∑ a ∈ s, f a) = ∑ a ∈ s, (f a : ℝ) := +theorem coe_sum (s : Finset ι) (f : ι → ℝ≥0) : ∑ i ∈ s, f i = ∑ i ∈ s, (f i : ℝ) := map_sum toRealHom _ _ -theorem _root_.Real.toNNReal_sum_of_nonneg {α} {s : Finset α} {f : α → ℝ} - (hf : ∀ a, a ∈ s → 0 ≤ f a) : +@[simp, norm_cast] +lemma coe_expect (s : Finset ι) (f : ι → ℝ≥0) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : ℝ) := + map_expect toRealHom .. + +theorem _root_.Real.toNNReal_sum_of_nonneg (hf : ∀ i ∈ s, 0 ≤ f i) : Real.toNNReal (∑ a ∈ s, f a) = ∑ a ∈ s, Real.toNNReal (f a) := by rw [← coe_inj, NNReal.coe_sum, Real.coe_toNNReal _ (Finset.sum_nonneg hf)] exact Finset.sum_congr rfl fun x hxs => by rw [Real.coe_toNNReal _ (hf x hxs)] @[simp, norm_cast] -theorem coe_prod {α} {s : Finset α} {f : α → ℝ≥0} : ↑(∏ a ∈ s, f a) = ∏ a ∈ s, (f a : ℝ) := +theorem coe_prod (s : Finset ι) (f : ι → ℝ≥0) : ↑(∏ a ∈ s, f a) = ∏ a ∈ s, (f a : ℝ) := map_prod toRealHom _ _ -theorem _root_.Real.toNNReal_prod_of_nonneg {α} {s : Finset α} {f : α → ℝ} - (hf : ∀ a, a ∈ s → 0 ≤ f a) : +theorem _root_.Real.toNNReal_prod_of_nonneg (hf : ∀ a, a ∈ s → 0 ≤ f a) : Real.toNNReal (∏ a ∈ s, f a) = ∏ a ∈ s, Real.toNNReal (f a) := by 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)] diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 5b895c2cdeb63..16325c2f3e0ab 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -3,16 +3,17 @@ 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.Data.Real.Star -import Mathlib.Algebra.Algebra.Basic +import Mathlib.Algebra.Module.Rat +import Mathlib.Algebra.Module.Submodule.Lattice import Mathlib.Algebra.Periodic +import Mathlib.Data.Real.Star import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Algebra.Order.Field -import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Algebra.Star +import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Instances.Int -import Mathlib.Topology.Order.Bornology import Mathlib.Topology.Metrizable.Basic +import Mathlib.Topology.Order.Bornology /-! # Topological properties of ℝ From 90a6b826db7d725b1e72b29d894fa45665d857e1 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Tue, 1 Oct 2024 18:15:49 +0000 Subject: [PATCH 091/174] feat(RingTheory/LaurentSeries): prove that Laurent Series over a field are complete (#16865) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR we provide a proof that the Laurent series over a field are complete. Co-authored-by: María Inés de Frutos-Fernández @mariainesdff --- Mathlib/Order/Filter/Ultrafilter.lean | 4 + Mathlib/RingTheory/LaurentSeries.lean | 181 +++++++++++++++++++++- Mathlib/Topology/UniformSpace/Basic.lean | 25 +++ Mathlib/Topology/UniformSpace/Cauchy.lean | 25 +++ 4 files changed, 232 insertions(+), 3 deletions(-) diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter.lean index 3d1ae9a87f5fc..6f238409f1d35 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter.lean @@ -465,4 +465,8 @@ theorem ofComapInfPrincipal_eq_of_map (h : m '' s ∈ g) : (ofComapInfPrincipal _ ≤ ↑g ⊓ (𝓟 <| m '' s) := inf_le_inf_right _ map_comap_le _ = ↑g := inf_of_le_left (le_principal_iff.mpr h) +theorem eq_of_le_pure {X : Type _} {α : Filter X} (hα : α.NeBot) {x y : X} + (hx : α ≤ pure x) (hy : α ≤ pure y) : x = y := + Filter.pure_injective (hα.le_pure_iff.mp hx ▸ hα.le_pure_iff.mp hy) + end Ultrafilter diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index e2ae574e5cae7..669be5c1d5cd1 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -11,6 +11,7 @@ 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 /-! # Laurent Series @@ -26,16 +27,23 @@ import Mathlib.RingTheory.Localization.FractionRing * Embedding of rational functions into Laurent series, provided as a coercion, utilizing the underlying `RatFunc.coeAlgHom`. * Study of the `X`-Adic valuation on the ring of Laurent series over a field +* In `LaurentSeries.uniformContinuous_coeff` we show that sending a Laurent series to its `d`th +coefficient is uniformly continuous, ensuring that it sends a Cauchy filter `ℱ` in `LaurentSeries K` +to a Cauchy filter in `K`: since this latter is given the discrete topology, this provides an +element `LaurentSeries.Cauchy.coeff ℱ d` in `K` that serves as `d`th coefficient of the Laurent +series to which the filter `ℱ` converges. ## Main Results * Basic properties of Hasse derivatives ### About the `X`-Adic valuation: * The (integral) valuation of a power series is the order of the first non-zero coefficient, see -`intValuation_le_iff_coeff_lt_eq_zero`. +`LaurentSeries.intValuation_le_iff_coeff_lt_eq_zero`. * The valuation of a Laurent series is the order of the first non-zero coefficient, see -`valuation_le_iff_coeff_lt_eq_zero`. +`LaurentSeries.valuation_le_iff_coeff_lt_eq_zero`. * Every Laurent series of valuation less than `(1 : ℤₘ₀)` comes from a power series, see -`val_le_one_iff_eq_coe`. +`LaurentSeries.val_le_one_iff_eq_coe`. +* The uniform space of `LaurentSeries` over a field is complete, formalized in the instance +`instLaurentSeriesComplete`. ## Implementation details * Since `LaurentSeries` is just an abbreviation of `HahnSeries ℤ _`, the definition of the @@ -464,6 +472,7 @@ end RatFunc namespace LaurentSeries + open IsDedekindDomain.HeightOneSpectrum PowerSeries RatFunc instance : Valued (LaurentSeries K) ℤₘ₀ := Valued.mk' (PowerSeries.idealX K).valuation @@ -609,3 +618,169 @@ theorem val_le_one_iff_eq_coe (f : LaurentSeries K) : Valued.v f ≤ (1 : ℤₘ end LaurentSeries end AdicValuation +namespace LaurentSeries +section Complete + +open Filter + +open scoped Multiplicative + +variable {K : Type*} [Field K] + +/- Sending a Laurent series to its `d`-th coefficient is uniformly continuous (independently of the + uniformity with which `K` is endowed). -/ +theorem uniformContinuous_coeff {uK : UniformSpace K} (d : ℤ) : + UniformContinuous fun f : LaurentSeries K ↦ f.coeff d := by + refine uniformContinuous_iff_eventually.mpr fun S hS ↦ eventually_iff_exists_mem.mpr ?_ + let γ : ℤₘ₀ˣ := Units.mk0 (↑(Multiplicative.ofAdd (-(d + 1)))) WithZero.coe_ne_zero + use {P | Valued.v (P.snd - P.fst) < ↑γ} + refine ⟨(Valued.hasBasis_uniformity (LaurentSeries K) ℤₘ₀).mem_of_mem (by tauto), fun P hP ↦ ?_⟩ + rw [eq_coeff_of_valuation_sub_lt K (le_of_lt hP) (lt_add_one _)] + exact mem_uniformity_of_eq hS rfl + +/-- Since extracting coefficients is uniformly continuous, every Cauchy filter in +`laurentSeries K` gives rise to a Cauchy filter in `K` for every `d : ℤ`, and such Cauchy filter +in `K` converges to a principal filter -/ +def Cauchy.coeff {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : ℤ → K := + let _ : UniformSpace K := ⊥ + fun d ↦ UniformSpace.DiscreteUnif.cauchyConst rfl <| hℱ.map (uniformContinuous_coeff d) + +theorem Cauchy.coeff_tendsto {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) (D : ℤ) : + Tendsto (fun f : LaurentSeries K ↦ f.coeff D) ℱ (𝓟 {coeff hℱ D}) := + let _ : UniformSpace K := ⊥ + le_of_eq <| UniformSpace.DiscreteUnif.eq_const_of_cauchy (by rfl) + (hℱ.map (uniformContinuous_coeff D)) ▸ (principal_singleton _).symm + +/- For every Cauchy filter of Laurent series, there is a `N` such that the `n`-th coefficient +vanishes for all `n ≤ N` and almost all series in the filter. This is an auxiliary lemma used +to construct the limit of the Cauchy filter as a Laurent series, ensuring that the support of the +limit is `PWO`. +The result is true also for more general Hahn Series indexed over a partially ordered group `Γ` +beyond the special case `Γ = ℤ`, that corresponds to Laurent Series: nevertheless the proof below +does not generalise, as it relies on the study of the `X`-adic valuation attached to the height-one +prime `X`, and this is peculiar to the one-variable setting. In the future we should prove this +result in full generality and deduce the case `Γ = ℤ` from that one.-/ +lemma Cauchy.exists_lb_eventual_support {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ᶠ f : LaurentSeries K in ℱ, ∀ n < N, f.coeff n = (0 : K) := by + let entourage : Set (LaurentSeries K × LaurentSeries K) := + {P : LaurentSeries K × LaurentSeries K | + Valued.v (P.snd - P.fst) < ((Multiplicative.ofAdd 0 : Multiplicative ℤ) : ℤₘ₀)} + let ζ := Units.mk0 (G₀ := ℤₘ₀) _ (WithZero.coe_ne_zero (a := (Multiplicative.ofAdd 0))) + obtain ⟨S, ⟨hS, ⟨T, ⟨hT, H⟩⟩⟩⟩ := mem_prod_iff.mp <| Filter.le_def.mp hℱ.2 entourage + <| (Valued.hasBasis_uniformity (LaurentSeries K) ℤₘ₀).mem_of_mem (i := ζ) (by tauto) + obtain ⟨f, hf⟩ := forall_mem_nonempty_iff_neBot.mpr hℱ.1 (S ∩ T) (inter_mem_iff.mpr ⟨hS, hT⟩) + obtain ⟨N, hN⟩ : ∃ N : ℤ, ∀ g : LaurentSeries K, + Valued.v (g - f) ≤ ↑(Multiplicative.ofAdd (0 : ℤ)) → ∀ n < N, g.coeff n = 0 := by + by_cases hf : f = 0 + · refine ⟨0, fun x hg ↦ ?_⟩ + rw [hf, sub_zero] at hg + exact (valuation_le_iff_coeff_lt_eq_zero K).mp hg + · refine ⟨min (f.2.isWF.min (HahnSeries.support_nonempty_iff.mpr hf)) 0 - 1, fun _ hg n hn ↦ ?_⟩ + rw [eq_coeff_of_valuation_sub_lt K hg (d := 0)] + · exact Function.nmem_support.mp fun h ↦ + f.2.isWF.not_lt_min (HahnSeries.support_nonempty_iff.mpr hf) h + <| lt_trans hn <| Int.sub_one_lt_iff.mpr <| min_le_left _ _ + exact lt_of_lt_of_le hn <| le_of_lt (Int.sub_one_lt_of_le <| min_le_right _ _) + use N + apply mem_of_superset (inter_mem hS hT) + intro g hg + have h_prod : (f, g) ∈ entourage := Set.prod_mono (Set.inter_subset_left (t := T)) + (Set.inter_subset_right (s := S)) |>.trans H <| Set.mem_prod.mpr ⟨hf, hg⟩ + exact hN g (le_of_lt h_prod) + +/- The support of `Cauchy.coeff` has a lower bound. -/ +theorem Cauchy.exists_lb_support {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ n, n < N → coeff hℱ n = 0 := by + let _ : UniformSpace K := ⊥ + obtain ⟨N, hN⟩ := exists_lb_eventual_support hℱ + refine ⟨N, fun n hn ↦ Ultrafilter.eq_of_le_pure (hℱ.map (uniformContinuous_coeff n)).1 + ((principal_singleton _).symm ▸ coeff_tendsto _ _) ?_⟩ + simp only [pure_zero, nonpos_iff] + apply Filter.mem_of_superset hN (fun _ ha ↦ ha _ hn) + +/- The support of `Cauchy.coeff` is bounded below -/ +theorem Cauchy.coeff_support_bddBelow {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + BddBelow (coeff hℱ).support := by + refine ⟨(exists_lb_support hℱ).choose, fun d hd ↦ ?_⟩ + by_contra hNd + exact hd ((exists_lb_support hℱ).choose_spec d (not_le.mp hNd)) + +/-- To any Cauchy filter ℱ of `LaurentSeries K`, we can attach a laurent series that is the limit +of the filter. Its `d`-th coefficient is defined as the limit of `Cauchy.coeff hℱ d`, which is +again Cauchy but valued in the discrete space `K`. That sufficiently negative coefficients vanish +follows from `Cauchy.coeff_support_bddBelow` -/ +def Cauchy.limit {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : LaurentSeries K := + HahnSeries.mk (coeff hℱ) <| Set.IsWF.isPWO (coeff_support_bddBelow _).wellFoundedOn_lt + +/- The following lemma shows that for every `d` smaller than the minimum between the integers +produced in `Cauchy.exists_lb_eventual_support` and `Cauchy.exists_lb_support`, for almost all +series in `ℱ` the `d`th coefficient coincides with the `d`th coefficient of `Cauchy.coeff hℱ`. -/ +theorem Cauchy.exists_lb_coeff_ne {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ᶠ f : LaurentSeries K in ℱ, ∀ d < N, coeff hℱ d = f.coeff d := by + obtain ⟨⟨N₁, hN₁⟩, ⟨N₂, hN₂⟩⟩ := exists_lb_eventual_support hℱ, exists_lb_support hℱ + refine ⟨min N₁ N₂, ℱ.3 hN₁ fun _ hf d hd ↦ ?_⟩ + rw [hf d (lt_of_lt_of_le hd (min_le_left _ _)), hN₂ d (lt_of_lt_of_le hd (min_le_right _ _))] + +/- Given a Cauchy filter `ℱ` in the Laurent Series and a bound `D`, for almost all series in the +filter the coefficients below `D` coincide with `Caucy.coeff hℱ`-/ +theorem Cauchy.coeff_eventually_equal {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) {D : ℤ} : + ∀ᶠ f : LaurentSeries K in ℱ, ∀ d, d < D → coeff hℱ d = f.coeff d := by + -- `φ` sends `d` to the set of Laurent Series having `d`th coefficient equal to `ℱ.coeff`. + let φ : ℤ → Set (LaurentSeries K) := fun d ↦ {f | coeff hℱ d = f.coeff d} + have intersec₁ : + (⋂ n ∈ Set.Iio D, φ n) ⊆ {x : LaurentSeries K | ∀ d : ℤ, d < D → coeff hℱ d = x.coeff d} := by + intro _ hf + simpa only [Set.mem_iInter] using hf + -- The goal is now to show that the intersection of all `φ d` (for `d < D`) is in `ℱ`. + let ℓ := (exists_lb_coeff_ne hℱ).choose + let N := max ℓ D + have intersec₂ : ⋂ n ∈ Set.Iio D, φ n ⊇ (⋂ n ∈ Set.Iio ℓ, φ n) ∩ (⋂ n ∈ Set.Icc ℓ N, φ n) := by + simp only [Set.mem_Iio, Set.mem_Icc, Set.subset_iInter_iff] + intro i hi x hx + simp only [Set.mem_inter_iff, Set.mem_iInter, and_imp] at hx + by_cases H : i < ℓ + exacts [hx.1 _ H, hx.2 _ (le_of_not_lt H) <| le_of_lt <| lt_max_of_lt_right hi] + suffices (⋂ n ∈ Set.Iio ℓ, φ n) ∩ (⋂ n ∈ Set.Icc ℓ N, φ n) ∈ ℱ by + exact ℱ.sets_of_superset this <| intersec₂.trans intersec₁ + /- To show that the intersection we have in sight is in `ℱ`, we use that it contains a double + intersection (an infinite and a finite one): by general properties of filters, we are reduced + to show that both terms are in `ℱ`, which is easy in light of their definition. -/ + · simp only [Set.mem_Iio, Set.mem_Ico, inter_mem_iff] + constructor + · have := (exists_lb_coeff_ne hℱ).choose_spec + rw [Filter.eventually_iff] at this + convert this + ext + simp only [Set.mem_iInter, Set.mem_setOf_eq]; rfl + · rw [biInter_mem (Set.finite_Icc ℓ N)] + intro _ _ + apply coeff_tendsto hℱ + simp only [principal_singleton, mem_pure]; rfl + + +open scoped Topology + +/- The main result showing that the Cauchy filter tends to the `Cauchy.limit`-/ +theorem Cauchy.eventually_mem_nhds {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) + {U : Set (LaurentSeries K)} (hU : U ∈ 𝓝 (Cauchy.limit hℱ)) : ∀ᶠ f in ℱ, f ∈ U := by + obtain ⟨γ, hU₁⟩ := Valued.mem_nhds.mp hU + suffices ∀ᶠ f in ℱ, f ∈ {y : LaurentSeries K | Valued.v (y - limit hℱ) < ↑γ} by + apply this.mono fun _ hf ↦ hU₁ hf + set D := -(Multiplicative.toAdd (WithZero.unzero γ.ne_zero) - 1) with hD₀ + have hD : ((Multiplicative.ofAdd (-D) : Multiplicative ℤ) : ℤₘ₀) < γ := by + rw [← WithZero.coe_unzero γ.ne_zero, WithZero.coe_lt_coe, hD₀, neg_neg, ofAdd_sub, + ofAdd_toAdd, div_lt_comm, div_self', ← ofAdd_zero, Multiplicative.ofAdd_lt] + exact zero_lt_one + apply coeff_eventually_equal hℱ |>.mono + intro _ hf + apply lt_of_le_of_lt (valuation_le_iff_coeff_lt_eq_zero K |>.mpr _) hD + intro n hn + rw [HahnSeries.sub_coeff, sub_eq_zero, hf n hn |>.symm]; rfl + +/- Laurent Series with coefficients in a field are complete w.r.t. the `X`-adic valuation -/ +instance instLaurentSeriesComplete : CompleteSpace (LaurentSeries K) := + ⟨fun hℱ ↦ ⟨Cauchy.limit hℱ, fun _ hS ↦ Cauchy.eventually_mem_nhds hℱ hS⟩⟩ + +end Complete + +end LaurentSeries diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index b57b69003857a..3e05609edfe68 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -136,6 +136,31 @@ theorem mem_idRel {a b : α} : (a, b) ∈ @idRel α ↔ a = b := theorem idRel_subset {s : Set (α × α)} : idRel ⊆ s ↔ ∀ a, (a, a) ∈ s := by simp [subset_def] +theorem eq_singleton_left_of_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, S = {x} := by + rcases hS, hT with ⟨⟨s, hs⟩, ⟨t, ht⟩⟩ + refine ⟨s, eq_singleton_iff_nonempty_unique_mem.mpr ⟨⟨s, hs⟩, fun x hx ↦ ?_⟩⟩ + rw [prod_subset_iff] at h_diag + replace hs := h_diag s hs t ht + replace hx := h_diag x hx t ht + simp only [idRel, mem_setOf_eq] at hx hs + rwa [← hs] at hx + +theorem eq_singleton_right_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, T = {x} := by + rw [Set.prod_subset_iff] at h_diag + replace h_diag := fun x hx y hy => (h_diag y hy x hx).symm + exact eq_singleton_left_of_prod_subset_idRel hT hS (prod_subset_iff.mpr h_diag) + +theorem eq_singleton_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, S = {x} ∧ T = {x} := by + obtain ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ := eq_singleton_left_of_prod_subset_idRel hS hT h_diag, + eq_singleton_right_prod_subset_idRel hS hT h_diag + refine ⟨x, ⟨hx, ?_⟩⟩ + rw [hy, Set.singleton_eq_singleton_iff] + exact (Set.prod_subset_iff.mp h_diag x (by simp only [hx, Set.mem_singleton]) y + (by simp only [hy, Set.mem_singleton])).symm + /-- The composition of relations -/ def compRel (r₁ r₂ : Set (α × α)) := { p : α × α | ∃ z : α, (p.1, z) ∈ r₁ ∧ (z, p.2) ∈ r₂ } diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 7907f4159bea3..a1a7590a0efa4 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -787,4 +787,29 @@ theorem secondCountable_of_separable [SeparableSpace α] : SecondCountableTopolo refine ⟨_, ⟨y, hys, k, rfl⟩, (hts k).subset hxy, fun z hz => ?_⟩ exact hUV (ball_subset_of_comp_subset (hk hxy) hUU' (hk hz)) +section DiscreteUniformity + +open Filter + +/-- A Cauchy filter in a discrete uniform space is contained in a principal filter-/ +theorem DiscreteUnif.cauchy_le_pure {X : Type _} {uX : UniformSpace X} + (hX : uX = ⊥) {α : Filter X} (hα : Cauchy α) : ∃ x : X, α = pure x := by + rcases hα with ⟨α_ne_bot, α_le⟩ + rw [hX, bot_uniformity, le_principal_iff, mem_prod_iff] at α_le + obtain ⟨S, ⟨hS, ⟨T, ⟨hT, H⟩⟩⟩⟩ := α_le + obtain ⟨x, rfl⟩ := eq_singleton_left_of_prod_subset_idRel (α_ne_bot.nonempty_of_mem hS) + (Filter.nonempty_of_mem hT) H + exact ⟨x, α_ne_bot.le_pure_iff.mp <| le_pure_iff.mpr hS⟩ + +/-- A constant to which a Cauchy filter in a discrete uniform space converges. -/ +noncomputable def DiscreteUnif.cauchyConst {X : Type _} {uX : UniformSpace X} + (hX : uX = ⊥) {α : Filter X} (hα : Cauchy α) : X := + (DiscreteUnif.cauchy_le_pure hX hα).choose + +theorem DiscreteUnif.eq_const_of_cauchy {X : Type _} {uX : UniformSpace X} (hX : uX = ⊥) + {α : Filter X} (hα : Cauchy α) : α = pure (DiscreteUnif.cauchyConst hX hα) := + (DiscreteUnif.cauchy_le_pure hX hα).choose_spec + +end DiscreteUniformity + end UniformSpace From 00acfa1c18743ab9581e039f90be4cf91fd36766 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 1 Oct 2024 18:56:51 +0000 Subject: [PATCH 092/174] feat(Algebra/TrivSqZeroExt): generalize some results to `Invertible` (#12125) I think this probably raises that a more general approach to handling partial inverses is needed soon, since it duplicates lots of the `Matrix` logic. --- Mathlib/Algebra/TrivSqZeroExt.lean | 94 +++++++++++++++++++++++++++--- 1 file changed, 85 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 0e976107f9917..4eb552d919fc2 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -690,7 +690,9 @@ section Inv variable {R : Type u} {M : Type v} variable [Neg M] [Inv R] [SMul Rᵐᵒᵖ M] [SMul R M] -/-- Inversion of the trivial-square-zero extension, sending $r + m$ to $r^{-1} - r^{-1}mr^{-1}$. -/ +/-- Inversion of the trivial-square-zero extension, sending $r + m$ to $r^{-1} - r^{-1}mr^{-1}$. + +Strictly this is only a _two_-sided inverse when the left and right actions associate. -/ instance instInv : Inv (tsze R M) := ⟨fun b => (b.1⁻¹, -(b.1⁻¹ •> b.2 <• b.1⁻¹))⟩ @@ -702,6 +704,75 @@ instance instInv : Inv (tsze R M) := end Inv +/-! This section is heavily inspired by analogous results about matrices. -/ +section Invertible +variable {R : Type u} {M : Type v} +variable [AddCommGroup M] [Semiring R] [Module Rᵐᵒᵖ M] [Module R M] + +/-- `x.fst : R` is invertible when `x : tzre R M` is. -/ +abbrev invertibleFstOfInvertible (x : tsze R M) [Invertible x] : Invertible x.fst where + invOf := (⅟x).fst + invOf_mul_self := by rw [← fst_mul, invOf_mul_self, fst_one] + mul_invOf_self := by rw [← fst_mul, mul_invOf_self, fst_one] + +theorem fst_invOf (x : tsze R M) [Invertible x] [Invertible x.fst] : (⅟x).fst = ⅟(x.fst) := by + letI := invertibleFstOfInvertible x + convert (rfl : _ = ⅟ x.fst) + +theorem mul_left_eq_one (r : R) (x : tsze R M) (h : r * x.fst = 1) : + (inl r + inr (-((r •> x.snd) <• r))) * x = 1 := by + ext <;> dsimp + · rw [add_zero, h] + · rw [add_zero, zero_add, smul_neg, op_smul_op_smul, h, op_one, one_smul, + add_neg_cancel] + +theorem mul_right_eq_one (x : tsze R M) (r : R) (h : x.fst * r = 1) : + x * (inl r + inr (-(r •> (x.snd <• r)))) = 1 := by + ext <;> dsimp + · rw [add_zero, h] + · rw [add_zero, zero_add, smul_neg, smul_smul, h, one_smul, neg_add_cancel] + +variable [SMulCommClass R Rᵐᵒᵖ M] + +/-- `x : tzre R M` is invertible when `x.fst : R` is. -/ +abbrev invertibleOfInvertibleFst (x : tsze R M) [Invertible x.fst] : Invertible x where + invOf := (⅟x.fst, -(⅟x.fst •> x.snd <• ⅟x.fst)) + invOf_mul_self := by + convert mul_left_eq_one _ _ (invOf_mul_self x.fst) + ext <;> simp + mul_invOf_self := by + convert mul_right_eq_one _ _ (mul_invOf_self x.fst) + ext <;> simp [smul_comm] + +theorem snd_invOf (x : tsze R M) [Invertible x] [Invertible x.fst] : + (⅟x).snd = -(⅟x.fst •> x.snd <• ⅟x.fst) := by + letI := invertibleOfInvertibleFst x + convert congr_arg (TrivSqZeroExt.snd (R := R) (M := M)) (_ : _ = ⅟ x) + convert rfl + +/-- Together `TrivSqZeroExt.detInvertibleOfInvertible` and `TrivSqZeroExt.invertibleOfDetInvertible` +form an equivalence, although both sides of the equiv are subsingleton anyway. -/ +@[simps] +def invertibleEquivInvertibleFst (x : tsze R M) : Invertible x ≃ Invertible x.fst where + toFun _ := invertibleFstOfInvertible x + invFun _ := invertibleOfInvertibleFst x + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- When lowered to a prop, `Matrix.invertibleEquivInvertibleFst` forms an `iff`. -/ +theorem isUnit_iff_isUnit_fst {x : tsze R M} : IsUnit x ↔ IsUnit x.fst := by + simp only [← nonempty_invertible_iff_isUnit, (invertibleEquivInvertibleFst x).nonempty_congr] + +@[simp] +theorem isUnit_inl_iff {r : R} : IsUnit (inl r : tsze R M) ↔ IsUnit r := by + rw [isUnit_iff_isUnit_fst, fst_inl] + +@[simp] +theorem isUnit_inr_iff {m : M} : IsUnit (inr m : tsze R M) ↔ Subsingleton R := by + simp_rw [isUnit_iff_isUnit_fst, fst_inr, isUnit_zero_iff, subsingleton_iff_zero_eq_one] + +end Invertible + section DivisionSemiring variable {R : Type u} {M : Type v} variable [DivisionSemiring R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] @@ -727,18 +798,19 @@ protected theorem inv_one : (1 : tsze R M)⁻¹ = (1 : tsze R M) := by rw [← inl_one, TrivSqZeroExt.inv_inl, inv_one] protected theorem inv_mul_cancel {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹ * x = 1 := by - ext - · rw [fst_mul, fst_inv, inv_mul_cancel₀ hx, fst_one] - · rw [snd_mul, snd_inv, snd_one, smul_neg, op_smul_op_smul, inv_mul_cancel₀ hx, op_one, one_smul, - fst_inv, add_neg_cancel] + convert mul_left_eq_one _ _ (_root_.inv_mul_cancel₀ hx) using 2 + ext <;> simp variable [SMulCommClass R Rᵐᵒᵖ M] +@[simp] theorem invOf_eq_inv (x : tsze R M) [Invertible x] : ⅟x = x⁻¹ := by + letI := invertibleFstOfInvertible x + ext <;> simp [fst_invOf, snd_invOf] + protected theorem mul_inv_cancel {x : tsze R M} (hx : fst x ≠ 0) : x * x⁻¹ = 1 := by - ext - · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel₀ hx] - · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel₀ hx, one_smul, - fst_inv, neg_add_cancel] + have : Invertible x.fst := Units.invertible (.mk0 _ hx) + have := invertibleOfInvertibleFst x + rw [← invOf_eq_inv, mul_invOf_self] protected theorem mul_inv_rev (a b : tsze R M) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by @@ -763,6 +835,10 @@ protected theorem inv_inv {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹⁻¹ = x := rw [fst_inv] apply inv_ne_zero hx +@[simp] +theorem isUnit_inv_iff {x : tsze R M} : IsUnit x⁻¹ ↔ IsUnit x := by + simp_rw [isUnit_iff_isUnit_fst, fst_inv, isUnit_iff_ne_zero, ne_eq, inv_eq_zero] + end DivisionSemiring section DivisionRing From 3a0cb01a5d950ecf9dc2c5edf69a3c9a2191b2a1 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 1 Oct 2024 19:58:38 +0000 Subject: [PATCH 093/174] chore: rename Polynomial.prod_roots_eq_coeff_zero_of_monic_of_split (#17337) --- Mathlib/Algebra/Polynomial/Splits.lean | 5 ++++- Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean | 2 +- Mathlib/RingTheory/Norm/Basic.lean | 2 +- 3 files changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 955c85eca5bf7..3b0f053ebb0ea 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -437,7 +437,7 @@ theorem aeval_root_derivative_of_splits [Algebra K L] [DecidableEq L] {P : K[X]} rw [eval_multiset_prod_X_sub_C_derivative hr] /-- If `P` is a monic polynomial that splits, then `coeff P 0` equals the product of the roots. -/ -theorem prod_roots_eq_coeff_zero_of_monic_of_split {P : K[X]} (hmo : P.Monic) +theorem prod_roots_eq_coeff_zero_of_monic_of_splits {P : K[X]} (hmo : P.Monic) (hP : P.Splits (RingHom.id K)) : coeff P 0 = (-1) ^ P.natDegree * P.roots.prod := by nth_rw 1 [eq_prod_roots_of_monic_of_splits_id hmo hP] rw [coeff_zero_eq_eval_zero, eval_multiset_prod, Multiset.map_map] @@ -449,6 +449,9 @@ theorem prod_roots_eq_coeff_zero_of_monic_of_split {P : K[X]} (hmo : P.Monic) rw [neg_eq_neg_one_mul] simp only [splits_iff_card_roots.1 hP, neg_mul, one_mul, Multiset.prod_map_neg] +@[deprecated (since := "2024-10-01")] +alias prod_roots_eq_coeff_zero_of_monic_of_split := prod_roots_eq_coeff_zero_of_monic_of_splits + /-- If `P` is a monic polynomial that splits, then `P.nextCoeff` equals the sum of the roots. -/ theorem sum_roots_eq_nextCoeff_of_monic_of_split {P : K[X]} (hmo : P.Monic) (hP : P.Splits (RingHom.id K)) : P.nextCoeff = -P.roots.sum := by diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean index ccd5c26310966..4eb330442505f 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean @@ -58,7 +58,7 @@ namespace Matrix theorem det_eq_prod_roots_charpoly_of_splits (hAps : A.charpoly.Splits (RingHom.id R)) : A.det = (Matrix.charpoly A).roots.prod := by rw [det_eq_sign_charpoly_coeff, ← charpoly_natDegree_eq_dim A, - Polynomial.prod_roots_eq_coeff_zero_of_monic_of_split A.charpoly_monic hAps, ← mul_assoc, + Polynomial.prod_roots_eq_coeff_zero_of_monic_of_splits A.charpoly_monic hAps, ← mul_assoc, ← pow_two, pow_right_comm, neg_one_sq, one_pow, one_mul] theorem trace_eq_sum_roots_charpoly_of_splits (hAps : A.charpoly.Splits (RingHom.id R)) : diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index 98235e2bf157f..142c31df3e1f0 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -74,7 +74,7 @@ theorem PowerBasis.norm_gen_eq_prod_roots [Algebra R F] (pb : PowerBasis R S) have := minpoly.monic pb.isIntegral_gen rw [PowerBasis.norm_gen_eq_coeff_zero_minpoly, ← pb.natDegree_minpoly, RingHom.map_mul, ← coeff_map, - prod_roots_eq_coeff_zero_of_monic_of_split (this.map _) ((splits_id_iff_splits _).2 hf), + prod_roots_eq_coeff_zero_of_monic_of_splits (this.map _) ((splits_id_iff_splits _).2 hf), this.natDegree_map, map_pow, ← mul_assoc, ← mul_pow] simp only [map_neg, _root_.map_one, neg_mul, neg_neg, one_pow, one_mul] From 96ac52ffd304f6da3498f0ba1e4e950e6f231b38 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Tue, 1 Oct 2024 20:15:55 +0000 Subject: [PATCH 094/174] chore(*.ModEq): reduce dependencies (#17154) Both `Nat.ModEq` and `Int.ModEq` have some spurious dependencies which can be removed by reasonable swaps of generic tactics/declarations for `Nat` and `Int` native versions (e.g. use `omega` more). This PR does that and makes a slouching attempt to hold back the import graph downstream of these. Co-authored-by: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> --- Mathlib/Algebra/CharP/Defs.lean | 5 +- Mathlib/Algebra/ModEq.lean | 1 + .../Algebra/Order/Group/Unbundled/Int.lean | 9 ++++ .../Combinatorics/Additive/FreimanHom.lean | 1 + Mathlib/Data/Int/ModEq.lean | 19 +++---- Mathlib/Data/Int/Order/Lemmas.lean | 11 ---- Mathlib/Data/Nat/ModEq.lean | 54 +++++++++---------- Mathlib/Data/ZMod/Defs.lean | 1 + Mathlib/Logic/Godel/GodelBetaFunction.lean | 3 +- Mathlib/Order/Filter/AtTopBot/ModEq.lean | 5 +- Mathlib/Tactic/ModCases.lean | 1 + 11 files changed, 54 insertions(+), 56 deletions(-) diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index d470a6086987e..34b448a97de44 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -11,6 +11,7 @@ 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 /-! # Characteristic of semirings @@ -102,10 +103,10 @@ lemma intCast_injOn_Ico [IsRightCancelAdd R] : InjOn (Int.cast : ℤ → R) (Ico 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, ← dvd_neg] + · rw [← neg_eq_zero, ← Int.cast_neg, ← Int.dvd_neg] lift -a to ℕ using neg_nonneg.mpr (le_of_lt h) with b rw [Int.cast_natCast, CharP.cast_eq_zero_iff R p, Int.natCast_dvd_natCast] - · simp only [Int.cast_zero, eq_self_iff_true, dvd_zero] + · simp only [Int.cast_zero, eq_self_iff_true, Int.dvd_zero] · lift a to ℕ using le_of_lt h with b rw [Int.cast_natCast, CharP.cast_eq_zero_iff R p, Int.natCast_dvd_natCast] diff --git a/Mathlib/Algebra/ModEq.lean b/Mathlib/Algebra/ModEq.lean index 8c4bcc0fd4def..a275aad2c7254 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.Order.Ring.Int import Mathlib.GroupTheory.QuotientGroup.Basic /-! diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean index 773e4ce5f46eb..96c6eed502379 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -82,6 +82,15 @@ theorem abs_le_one_iff {a : ℤ} : |a| ≤ 1 ↔ a = 0 ∨ a = 1 ∨ a = -1 := b theorem one_le_abs {z : ℤ} (h₀ : z ≠ 0) : 1 ≤ |z| := add_one_le_iff.mpr (abs_pos.mpr h₀) +lemma eq_zero_of_abs_lt_dvd {m x : ℤ} (h1 : m ∣ x) (h2 : |x| < m) : x = 0 := by + by_contra h + have := Int.natAbs_le_of_dvd_ne_zero h1 h + rw [Int.abs_eq_natAbs] at h2 + omega + +lemma abs_sub_lt_of_lt_lt {m a b : ℕ} (ha : a < m) (hb : b < m) : |(b : ℤ) - a| < m := by + rw [abs_lt]; omega + /-! #### `/` -/ theorem ediv_eq_zero_of_lt_abs {a b : ℤ} (H1 : 0 ≤ a) (H2 : a < |b|) : a / b = 0 := diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index dee5a50bf940d..01cf0ef4b01c3 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.ZMod.Defs /-! diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index 2f973a2e81d79..d171639d57b4d 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Abel -import Mathlib.Tactic.GCongr.CoreAttrs /-! @@ -92,8 +90,7 @@ theorem mod_modEq (a n) : a % n ≡ a [ZMOD n] := @[simp] theorem neg_modEq_neg : -a ≡ -b [ZMOD n] ↔ a ≡ b [ZMOD n] := by --- Porting note: Restore old proof once #3309 is through - simp [-sub_neg_eq_add, neg_sub_neg, modEq_iff_dvd, dvd_sub_comm] + simp only [modEq_iff_dvd, (by omega : -b - -a = -(b - a)), Int.dvd_neg] @[simp] theorem modEq_neg : a ≡ b [ZMOD -n] ↔ a ≡ b [ZMOD n] := by simp [modEq_iff_dvd] @@ -105,9 +102,9 @@ protected theorem of_dvd (d : m ∣ n) (h : a ≡ b [ZMOD n]) : a ≡ b [ZMOD m] protected theorem mul_left' (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD c * n] := by obtain hc | rfl | hc := lt_trichotomy c 0 - · rw [← neg_modEq_neg, ← modEq_neg, ← neg_mul, ← neg_mul, ← neg_mul] + · rw [← neg_modEq_neg, ← modEq_neg, ← Int.neg_mul, ← Int.neg_mul, ← Int.neg_mul] simp only [ModEq, mul_emod_mul_of_pos _ _ (neg_pos.2 hc), h.eq] - · simp only [zero_mul, ModEq.rfl] + · simp only [Int.zero_mul, ModEq.rfl] · simp only [ModEq, mul_emod_mul_of_pos _ _ hc, h.eq] protected theorem mul_right' (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n * c] := by @@ -115,7 +112,7 @@ protected theorem mul_right' (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n * @[gcongr] protected theorem add (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a + c ≡ b + d [ZMOD n] := - modEq_iff_dvd.2 <| by convert dvd_add h₁.dvd h₂.dvd using 1; abel + modEq_iff_dvd.2 <| by convert Int.dvd_add h₁.dvd h₂.dvd using 1; omega @[gcongr] protected theorem add_left (c : ℤ) (h : a ≡ b [ZMOD n]) : c + a ≡ c + b [ZMOD n] := ModEq.rfl.add h @@ -125,10 +122,10 @@ protected theorem add (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a + protected theorem add_left_cancel (h₁ : a ≡ b [ZMOD n]) (h₂ : a + c ≡ b + d [ZMOD n]) : c ≡ d [ZMOD n] := - have : d - c = b + d - (a + c) - (b - a) := by abel + have : d - c = b + d - (a + c) - (b - a) := by omega modEq_iff_dvd.2 <| by rw [this] - exact dvd_sub h₂.dvd h₁.dvd + exact Int.dvd_sub h₂.dvd h₁.dvd protected theorem add_left_cancel' (c : ℤ) (h : c + a ≡ c + b [ZMOD n]) : a ≡ b [ZMOD n] := ModEq.rfl.add_left_cancel h @@ -183,7 +180,7 @@ theorem cancel_right_div_gcd (hm : 0 < m) (h : a * c ≡ b * c [ZMOD m]) : rw [modEq_iff_dvd] at h ⊢ -- Porting note: removed `show` due to leanprover-community/mathlib4#3305 refine Int.dvd_of_dvd_mul_right_of_gcd_one (?_ : m / d ∣ c / d * (b - a)) ?_ - · rw [mul_comm, ← Int.mul_ediv_assoc (b - a) gcd_dvd_right, sub_mul] + · rw [mul_comm, ← Int.mul_ediv_assoc (b - a) gcd_dvd_right, Int.sub_mul] exact Int.ediv_dvd_ediv gcd_dvd_left h · rw [gcd_div gcd_dvd_left gcd_dvd_right, natAbs_ofNat, Nat.div_self (gcd_pos_of_ne_zero_left c hm.ne')] @@ -233,7 +230,7 @@ theorem modEq_add_fac {a b n : ℤ} (c : ℤ) (ha : a ≡ b [ZMOD n]) : a + n * _ ≡ b [ZMOD n] := by rw [add_zero] theorem modEq_sub_fac {a b n : ℤ} (c : ℤ) (ha : a ≡ b [ZMOD n]) : a - n * c ≡ b [ZMOD n] := by - convert Int.modEq_add_fac (-c) ha using 1; rw [mul_neg, sub_eq_add_neg] + convert Int.modEq_add_fac (-c) ha using 1; rw [Int.mul_neg, sub_eq_add_neg] theorem modEq_add_fac_self {a t n : ℤ} : a + n * t ≡ a [ZMOD n] := modEq_add_fac _ ModEq.rfl diff --git a/Mathlib/Data/Int/Order/Lemmas.lean b/Mathlib/Data/Int/Order/Lemmas.lean index e7bc9e148fc42..d664a5b284aca 100644 --- a/Mathlib/Data/Int/Order/Lemmas.lean +++ b/Mathlib/Data/Int/Order/Lemmas.lean @@ -34,15 +34,4 @@ theorem natAbs_le_iff_mul_self_le {a b : ℤ} : a.natAbs ≤ b.natAbs ↔ a * a rw [← abs_le_iff_mul_self_le, abs_eq_natAbs, abs_eq_natAbs] exact Int.ofNat_le.symm -/-! ### units -/ - - -theorem eq_zero_of_abs_lt_dvd {m x : ℤ} (h1 : m ∣ x) (h2 : |x| < m) : x = 0 := by - obtain rfl | hm := eq_or_ne m 0 - · exact Int.zero_dvd.1 h1 - rcases h1 with ⟨d, rfl⟩ - apply mul_eq_zero_of_right - rw [← abs_lt_one_iff, ← mul_lt_iff_lt_one_right (abs_pos.mpr hm), ← abs_mul] - exact lt_of_lt_of_le h2 (le_abs_self m) - end Int diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index 5877782fc1102..29ff2fed4da9d 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -3,10 +3,8 @@ 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.Algebra.Ring.Regular +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Data.Int.GCD -import Mathlib.Data.Int.Order.Lemmas -import Mathlib.Tactic.NormNum.Basic /-! # Congruences modulo a natural number @@ -24,6 +22,7 @@ and proves basic properties about it such as the Chinese Remainder Theorem ModEq, congruence, mod, MOD, modulo -/ +assert_not_exists OrderedAddCommMonoid assert_not_exists Function.support namespace Nat @@ -37,8 +36,8 @@ notation:50 a " ≡ " b " [MOD " n "]" => ModEq n a b variable {m n a b c d : ℕ} --- Porting note: This instance should be derivable automatically -instance : Decidable (ModEq n a b) := decEq (a % n) (b % n) +-- Since `ModEq` is semi-reducible, we need to provide the decidable instance manually +instance : Decidable (ModEq n a b) := inferInstanceAs <| Decidable (a % n = b % n) namespace ModEq @@ -91,7 +90,7 @@ theorem mod_modEq (a n) : a % n ≡ a [MOD n] := namespace ModEq lemma of_dvd (d : m ∣ n) (h : a ≡ b [MOD n]) : a ≡ b [MOD m] := - modEq_of_dvd <| d.natCast.trans h.dvd + modEq_of_dvd <| Int.ofNat_dvd.mpr d |>.trans h.dvd protected theorem mul_left' (c : ℕ) (h : a ≡ b [MOD n]) : c * a ≡ c * b [MOD c * n] := by unfold ModEq at *; rw [mul_mod_mul_left, mul_mod_mul_left, h] @@ -122,7 +121,7 @@ protected theorem pow (m : ℕ) (h : a ≡ b [MOD n]) : a ^ m ≡ b ^ m [MOD n] @[gcongr] protected theorem add (h₁ : a ≡ b [MOD n]) (h₂ : c ≡ d [MOD n]) : a + c ≡ b + d [MOD n] := by rw [modEq_iff_dvd, Int.ofNat_add, Int.ofNat_add, add_sub_add_comm] - exact dvd_add h₁.dvd h₂.dvd + exact Int.dvd_add h₁.dvd h₂.dvd @[gcongr] protected theorem add_left (c : ℕ) (h : a ≡ b [MOD n]) : c + a ≡ c + b [MOD n] := @@ -136,7 +135,7 @@ protected theorem add_left_cancel (h₁ : a ≡ b [MOD n]) (h₂ : a + c ≡ b + c ≡ d [MOD n] := by simp only [modEq_iff_dvd, Int.ofNat_add] at * rw [add_sub_add_comm] at h₂ - convert _root_.dvd_sub h₂ h₁ using 1 + convert Int.dvd_sub h₂ h₁ using 1 rw [add_sub_cancel_left] protected theorem add_left_cancel' (c : ℕ) (h : c + a ≡ c + b [MOD n]) : a ≡ b [MOD n] := @@ -155,7 +154,8 @@ protected theorem add_right_cancel' (c : ℕ) (h : a + c ≡ b + c [MOD n]) : a For cancelling left multiplication in the modulus, see `Nat.ModEq.of_mul_left`. -/ protected theorem mul_left_cancel' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] → a ≡ b [MOD m] := by - simp [modEq_iff_dvd, ← mul_sub, mul_dvd_mul_iff_left (by simp [hc] : (c : ℤ) ≠ 0)] + simp only [modEq_iff_dvd, Int.natCast_mul, ← Int.mul_sub] + exact fun h => (Int.dvd_of_mul_dvd_mul_left (Int.ofNat_ne_zero.mpr hc) h) protected theorem mul_left_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] ↔ a ≡ b [MOD m] := @@ -166,7 +166,8 @@ protected theorem mul_left_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : For cancelling right multiplication in the modulus, see `Nat.ModEq.of_mul_right`. -/ protected theorem mul_right_cancel' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] → a ≡ b [MOD m] := by - simp [modEq_iff_dvd, ← sub_mul, mul_dvd_mul_iff_right (by simp [hc] : (c : ℤ) ≠ 0)] + simp only [modEq_iff_dvd, Int.natCast_mul, ← Int.sub_mul] + exact fun h => (Int.dvd_of_mul_dvd_mul_right (Int.ofNat_ne_zero.mpr hc) h) protected theorem mul_right_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] ↔ a ≡ b [MOD m] := @@ -204,10 +205,10 @@ namespace ModEq theorem le_of_lt_add (h1 : a ≡ b [MOD m]) (h2 : a < b + m) : a ≤ b := (le_total a b).elim id fun h3 => Nat.le_of_sub_eq_zero - (eq_zero_of_dvd_of_lt ((modEq_iff_dvd' h3).mp h1.symm) ((tsub_lt_iff_left h3).mpr h2)) + (eq_zero_of_dvd_of_lt ((modEq_iff_dvd' h3).mp h1.symm) (by omega)) theorem add_le_of_lt (h1 : a ≡ b [MOD m]) (h2 : a < b) : a + m ≤ b := - le_of_lt_add (add_modEq_right.trans h1) (add_lt_add_right h2 m) + le_of_lt_add (add_modEq_right.trans h1) (by omega) theorem dvd_iff (h : a ≡ b [MOD m]) (hdm : d ∣ m) : d ∣ a ↔ d ∣ b := by simp only [← modEq_zero_iff_dvd] @@ -227,9 +228,7 @@ lemma eq_of_abs_lt (h : a ≡ b [MOD m]) (h2 : |(b : ℤ) - a| < m) : a = b := b exact Int.eq_zero_of_abs_lt_dvd h.dvd h2 lemma eq_of_lt_of_lt (h : a ≡ b [MOD m]) (ha : a < m) (hb : b < m) : a = b := - h.eq_of_abs_lt <| abs_sub_lt_iff.2 - ⟨(sub_le_self _ <| Int.natCast_nonneg _).trans_lt <| Int.ofNat_lt.2 hb, - (sub_le_self _ <| Int.natCast_nonneg _).trans_lt <| Int.ofNat_lt.2 ha⟩ + h.eq_of_abs_lt <| Int.abs_sub_lt_of_lt_lt ha hb /-- To cancel a common factor `c` from a `ModEq` we must divide the modulus `m` by `gcd m c` -/ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b [MOD m / gcd m c] := by @@ -241,7 +240,7 @@ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b · show (m / d : ℤ) ∣ c / d * (b - a) rw [mul_comm, ← Int.mul_ediv_assoc (b - a) (Int.natCast_dvd_natCast.mpr hcd), mul_comm] apply Int.ediv_dvd_ediv (Int.natCast_dvd_natCast.mpr hmd) - rw [mul_sub] + rw [Int.mul_sub] exact modEq_iff_dvd.mp h · show Int.gcd (m / d) (c / d) = 1 simp only [← Int.natCast_div, Int.gcd_natCast_natCast (m / d) (c / d), gcd_div hmd hcd, @@ -299,18 +298,18 @@ def chineseRemainder' (h : a ≡ b [MOD gcd n m]) : { k // k ≡ a [MOD n] ∧ k have hcoedvd : ∀ t, (gcd n m : ℤ) ∣ t * (b - a) := fun t => h.dvd.mul_left _ have := gcd_eq_gcd_ab n m constructor <;> rw [Int.emod_def, ← sub_add] <;> - refine dvd_add ?_ (dvd_mul_of_dvd_left ?_ _) <;> + refine Int.dvd_add ?_ (dvd_mul_of_dvd_left ?_ _) <;> try norm_cast · rw [← sub_eq_iff_eq_add'] at this - rw [← this, sub_mul, ← add_sub_assoc, add_comm, add_sub_assoc, ← mul_sub, + rw [← this, Int.sub_mul, ← add_sub_assoc, add_comm, add_sub_assoc, ← Int.mul_sub, Int.add_ediv_of_dvd_left, Int.mul_ediv_cancel_left _ hnonzero, - Int.mul_ediv_assoc _ h.dvd, ← sub_sub, sub_self, zero_sub, dvd_neg, mul_assoc] + Int.mul_ediv_assoc _ h.dvd, ← sub_sub, sub_self, zero_sub, Int.dvd_neg, mul_assoc] · exact dvd_mul_right _ _ norm_cast exact dvd_mul_right _ _ · exact dvd_lcm_left n m · rw [← sub_eq_iff_eq_add] at this - rw [← this, sub_mul, sub_add, ← mul_sub, Int.sub_ediv_of_dvd, + rw [← this, Int.sub_mul, sub_add, ← Int.mul_sub, Int.sub_ediv_of_dvd, Int.mul_ediv_cancel_left _ hnonzero, Int.mul_ediv_assoc _ h.dvd, ← sub_add, sub_self, zero_add, mul_assoc] · exact dvd_mul_right _ _ @@ -407,7 +406,7 @@ protected theorem add_div_of_dvd_right {a b c : ℕ} (hca : c ∣ a) : (a + b) / add_div_eq_of_add_mod_lt (by rw [Nat.mod_eq_zero_of_dvd hca, zero_add] - exact Nat.mod_lt _ (pos_iff_ne_zero.mpr h)) + exact Nat.mod_lt _ (zero_lt_of_ne_zero h)) protected theorem add_div_of_dvd_left {a b c : ℕ} (hca : c ∣ b) : (a + b) / c = a / c + b / c := by rwa [add_comm, Nat.add_div_of_dvd_right, add_comm] @@ -430,27 +429,24 @@ theorem odd_mul_odd {n m : ℕ} : n % 2 = 1 → m % 2 = 1 → n * m % 2 = 1 := b theorem odd_mul_odd_div_two {m n : ℕ} (hm1 : m % 2 = 1) (hn1 : n % 2 = 1) : m * n / 2 = m * (n / 2) + m / 2 := - have hm0 : 0 < m := Nat.pos_of_ne_zero fun h => by simp_all have hn0 : 0 < n := Nat.pos_of_ne_zero fun h => by simp_all mul_right_injective₀ two_ne_zero <| by dsimp rw [mul_add, two_mul_odd_div_two hm1, mul_left_comm, two_mul_odd_div_two hn1, - two_mul_odd_div_two (Nat.odd_mul_odd hm1 hn1), mul_tsub, mul_one, ← - add_tsub_assoc_of_le (succ_le_of_lt hm0), - tsub_add_cancel_of_le (le_mul_of_one_le_right (Nat.zero_le _) hn0)] + two_mul_odd_div_two (Nat.odd_mul_odd hm1 hn1), Nat.mul_sub, mul_one, ← + Nat.add_sub_assoc (by omega), Nat.sub_add_cancel (Nat.le_mul_of_pos_right m hn0)] theorem odd_of_mod_four_eq_one {n : ℕ} : n % 4 = 1 → n % 2 = 1 := by - simpa [ModEq, show 2 * 2 = 4 by norm_num] using @ModEq.of_mul_left 2 n 1 2 + simpa [ModEq] using @ModEq.of_mul_left 2 n 1 2 theorem odd_of_mod_four_eq_three {n : ℕ} : n % 4 = 3 → n % 2 = 1 := by - simpa [ModEq, show 2 * 2 = 4 by norm_num, show 3 % 4 = 3 by norm_num] using - @ModEq.of_mul_left 2 n 3 2 + simpa [ModEq] using @ModEq.of_mul_left 2 n 3 2 /-- A natural number is odd iff it has residue `1` or `3` mod `4`-/ theorem odd_mod_four_iff {n : ℕ} : n % 2 = 1 ↔ n % 4 = 1 ∨ n % 4 = 3 := have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := by decide ⟨fun hn => - help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, + help (n % 4) (mod_lt n (by omega)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, fun h => Or.elim h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩ lemma mod_eq_of_modEq {a b n} (h : a ≡ b [MOD n]) (hb : b < n) : a % n = b := diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 324dec3252219..6977197f083cc 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -5,6 +5,7 @@ Authors: Eric Rodriguez -/ import Mathlib.Algebra.Group.Fin.Basic import Mathlib.Algebra.NeZero +import Mathlib.Algebra.Ring.Int import Mathlib.Data.Nat.ModEq import Mathlib.Data.Fintype.Card diff --git a/Mathlib/Logic/Godel/GodelBetaFunction.lean b/Mathlib/Logic/Godel/GodelBetaFunction.lean index 4fe1e162e2c5a..52ef094f4d92b 100644 --- a/Mathlib/Logic/Godel/GodelBetaFunction.lean +++ b/Mathlib/Logic/Godel/GodelBetaFunction.lean @@ -76,8 +76,7 @@ private lemma pairwise_coprime_coprimes (a : Fin m → ℕ) : Pairwise (Coprime have hja : j < supOfSeq a := lt_of_lt_of_le j.prop (le_step (le_max_left _ _)) exact coprime_mul_succ (Nat.succ_le_succ <| le_of_lt ltij) - (Nat.dvd_factorial - (by simp [Nat.succ_sub_succ, ltij]) + (Nat.dvd_factorial (by omega) (by simpa only [Nat.succ_sub_succ] using le_of_lt (lt_of_le_of_lt (sub_le j i) hja))) /-- Gödel's Beta Function. This is similar to `(Encodable.decodeList).get i`, but it is easier to diff --git a/Mathlib/Order/Filter/AtTopBot/ModEq.lean b/Mathlib/Order/Filter/AtTopBot/ModEq.lean index f7974d0c8d968..d45c17cc0d74b 100644 --- a/Mathlib/Order/Filter/AtTopBot/ModEq.lean +++ b/Mathlib/Order/Filter/AtTopBot/ModEq.lean @@ -3,7 +3,10 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Order.Ring.Abs +import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int import Mathlib.Data.Nat.ModEq import Mathlib.Order.Filter.AtTopBot.Monoid diff --git a/Mathlib/Tactic/ModCases.lean b/Mathlib/Tactic/ModCases.lean index 9846cca299457..fd1a5afdd358f 100644 --- a/Mathlib/Tactic/ModCases.lean +++ b/Mathlib/Tactic/ModCases.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth -/ import Mathlib.Data.Int.ModEq +import Mathlib.Tactic.HaveI /-! # `mod_cases` tactic From ebdf9cca3c351868ae7c2786096256474715e5e2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 20:15:56 +0000 Subject: [PATCH 095/174] feat(ENNReal): Convenience lemmas (#17277) Add a few convenience lemmas using the equality `a = b + c` to deduce that `b` is finite from the fact that `a` is. As shown by the golfs, this is useful. Also deprecate `sub_eq_of_add_eq`, which is the "annoyingly swap things around" version of `eq_sub_of_add_eq`/`sub_eq_of_eq_add`. From LeanAPAP --- Archive/Wiedijk100Theorems/BallotProblem.lean | 2 +- .../Order/BigOperators/Group/Finset.lean | 10 ++++++++ .../Algebra/Order/Monoid/Unbundled/Basic.lean | 18 +++++++++++++- Mathlib/Algebra/Order/Sub/Defs.lean | 22 +++++++++++++++++ Mathlib/Data/ENNReal/Inv.lean | 4 +--- Mathlib/Data/ENNReal/Operations.lean | 24 ++++++++++++++++++- Mathlib/Data/ENNReal/Real.lean | 2 +- .../Measure/Lebesgue/EqHaar.lean | 6 ++--- .../MeasureTheory/Measure/MeasureSpace.lean | 2 +- 9 files changed, 78 insertions(+), 12 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 5d22c4ad829b2..8fbd5b2ecf7d0 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -235,7 +235,7 @@ theorem first_vote_neg (p q : ℕ) (h : 0 < p + q) : have := condCount_compl {l : List ℤ | l.headI = 1}ᶜ (countedSequence_finite p q) (countedSequence_nonempty p q) rw [compl_compl, first_vote_pos _ _ h] at this - rw [← ENNReal.sub_eq_of_add_eq _ this, ENNReal.eq_div_iff, ENNReal.mul_sub, mul_one, + rw [ENNReal.eq_sub_of_add_eq _ this, ENNReal.eq_div_iff, ENNReal.mul_sub, mul_one, ENNReal.mul_div_cancel', ENNReal.add_sub_cancel_left] all_goals simp_all [ENNReal.div_eq_top] diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 2f1547de189c6..c7a0670ed2b05 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -221,6 +221,16 @@ theorem abs_sum_of_nonneg' {G : Type*} [LinearOrderedAddCommGroup G] {f : ι → (hf : ∀ i, 0 ≤ f i) : |∑ i ∈ s, f i| = ∑ i ∈ s, f i := by rw [abs_of_nonneg (Finset.sum_nonneg' hf)] +section CommMonoid +variable [CommMonoid α] [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {s : Finset ι} {f : ι → α} + +@[to_additive (attr := simp)] +lemma mulLECancellable_prod : + MulLECancellable (∏ i ∈ s, f i) ↔ ∀ ⦃i⦄, i ∈ s → MulLECancellable (f i) := by + induction' s using Finset.cons_induction with i s hi ih <;> simp [*] + +end CommMonoid + section Pigeonhole variable [DecidableEq β] diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index d6509bd4ccb64..f33502550e3da 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -1289,7 +1289,7 @@ theorem Contravariant.MulLECancellable [Mul α] [LE α] [ContravariantClass α MulLECancellable a := fun _ _ => le_of_mul_le_mul_left' -@[to_additive] +@[to_additive (attr := simp)] theorem mulLECancellable_one [Monoid α] [LE α] : MulLECancellable (1 : α) := fun a b => by simpa only [one_mul] using id @@ -1350,4 +1350,20 @@ protected theorem mul_le_iff_le_one_left [MulOneClass α] [i : @Std.Commutative [CovariantClass α α (· * ·) (· ≤ ·)] {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 : α} + (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 : α} + (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 : α} : + 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/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 08e4b08db1ebc..55d910166545c 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -243,17 +243,39 @@ theorem tsub_right_comm : a - b - c = a - c - b := by namespace AddLECancellable +/-- See `AddLECancellable.tsub_eq_of_eq_add'` for a version assuming that `a = c + b` itself is +cancellable rather than `b`. -/ protected theorem tsub_eq_of_eq_add (hb : AddLECancellable b) (h : a = c + b) : a - b = c := le_antisymm (tsub_le_iff_right.mpr h.le) <| by rw [h] exact hb.le_add_tsub +/-- 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) + (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 +cancellable rather than `c`. -/ protected theorem eq_tsub_of_add_eq (hc : AddLECancellable c) (h : a + c = b) : a = b - c := (hc.tsub_eq_of_eq_add h.symm).symm +/-- 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) + (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 +cancellable rather than `b`. -/ protected theorem tsub_eq_of_eq_add_rev (hb : AddLECancellable b) (h : a = b + c) : a - b = c := hb.tsub_eq_of_eq_add <| by rw [add_comm, h] +/-- 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 α α (· + ·) (· ≤ ·)] + (ha : AddLECancellable a) (h : a = b + c) : a - b = c := + ha.tsub_eq_of_eq_add' <| by rw [add_comm, h] + @[simp] protected theorem add_tsub_cancel_right (hb : AddLECancellable b) : a + b - b = a := hb.tsub_eq_of_eq_add <| by rw [add_comm] diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index d0e3217ac181c..1be528c876575 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -438,9 +438,7 @@ protected theorem half_lt_self (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := by protected theorem half_le_self : a / 2 ≤ a := le_add_self.trans_eq <| ENNReal.add_halves _ -theorem sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := by - lift a to ℝ≥0 using h - exact sub_eq_of_add_eq (mul_ne_top coe_ne_top <| by simp) (ENNReal.add_halves a) +theorem sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := ENNReal.sub_eq_of_eq_add' h a.add_halves.symm @[simp] theorem one_sub_inv_two : (1 : ℝ≥0∞) - 2⁻¹ = 2⁻¹ := by diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 9f0b1dbbbddc3..bf0335bc473f4 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -314,15 +314,37 @@ theorem natCast_sub (m n : ℕ) : ↑(m - n) = (m - n : ℝ≥0∞) := by @[deprecated (since := "2024-04-17")] alias nat_cast_sub := natCast_sub +/-- See `ENNReal.sub_eq_of_eq_add'` for a version assuming that `a = c + b` itself is finite rather +than `b`. -/ protected theorem sub_eq_of_eq_add (hb : b ≠ ∞) : a = c + b → a - b = c := (cancel_of_ne hb).tsub_eq_of_eq_add +/-- Weaker version of `ENNReal.sub_eq_of_eq_add` assuming that `a = c + b` itself is finite rather +han `b`. -/ +protected lemma sub_eq_of_eq_add' (ha : a ≠ ∞) : a = c + b → a - b = c := + (cancel_of_ne ha).tsub_eq_of_eq_add' + +/-- See `ENNReal.eq_sub_of_add_eq'` for a version assuming that `b = a + c` itself is finite rather +than `c`. -/ protected theorem eq_sub_of_add_eq (hc : c ≠ ∞) : a + c = b → a = b - c := (cancel_of_ne hc).eq_tsub_of_add_eq +/-- Weaker version of `ENNReal.eq_sub_of_add_eq` assuming that `b = a + c` itself is finite rather +than `c`. -/ +protected lemma eq_sub_of_add_eq' (hb : b ≠ ∞) : a + c = b → a = b - c := + (cancel_of_ne hb).eq_tsub_of_add_eq' + +/-- See `ENNReal.sub_eq_of_eq_add_rev'` for a version assuming that `a = b + c` itself is finite +rather than `b`. -/ protected theorem sub_eq_of_eq_add_rev (hb : b ≠ ∞) : a = b + c → a - b = c := (cancel_of_ne hb).tsub_eq_of_eq_add_rev +/-- Weaker version of `ENNReal.sub_eq_of_eq_add_rev` assuming that `a = b + c` itself is finite +rather than `b`. -/ +protected lemma sub_eq_of_eq_add_rev' (ha : a ≠ ∞) : a = b + c → a - b = c := + (cancel_of_ne ha).tsub_eq_of_eq_add_rev' + +@[deprecated ENNReal.sub_eq_of_eq_add (since := "2024-09-30")] theorem sub_eq_of_add_eq (hb : b ≠ ∞) (hc : a + b = c) : c - b = a := ENNReal.sub_eq_of_eq_add hb hc.symm @@ -338,7 +360,7 @@ protected theorem sub_add_eq_add_sub (hab : b ≤ a) (b_ne_top : b ≠ ∞) : a - b + c = a + c - b := by by_cases c_top : c = ∞ · simpa [c_top] using ENNReal.eq_sub_of_add_eq b_ne_top rfl - refine (sub_eq_of_add_eq b_ne_top ?_).symm + refine ENNReal.eq_sub_of_add_eq b_ne_top ?_ simp only [add_assoc, add_comm c b] simpa only [← add_assoc] using (add_left_inj c_top).mpr <| tsub_add_cancel_of_le hab diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 4b6df926ef130..972d93ec15a0c 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -235,7 +235,7 @@ lemma ofNat_le_ofReal {n : ℕ} [n.AtLeastTwo] {p : ℝ} : no_index (OfNat.ofNat n) ≤ ENNReal.ofReal p ↔ OfNat.ofNat n ≤ p := natCast_le_ofReal (NeZero.ne n) -@[simp] +@[simp, norm_cast] lemma ofReal_le_natCast {r : ℝ} {n : ℕ} : ENNReal.ofReal r ≤ n ↔ r ≤ n := coe_le_coe.trans Real.toNNReal_le_natCast diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 4f9ff4c003ac1..54c09f7095a7a 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -745,10 +745,8 @@ theorem tendsto_addHaar_inter_smul_one_of_density_one_aux (s : Set E) (hs : Meas rw [← ENNReal.sub_mul]; swap · simp only [uzero, ENNReal.inv_eq_top, imp_true_iff, Ne, not_false_iff] congr 1 - apply - ENNReal.sub_eq_of_add_eq (ne_top_of_le_ne_top utop (measure_mono inter_subset_right)) - rw [inter_comm _ u, inter_comm _ u] - exact measure_inter_add_diff u vmeas + rw [inter_comm _ u, inter_comm _ u, eq_comm] + exact ENNReal.eq_sub_of_add_eq' utop (measure_inter_add_diff u vmeas) have L : Tendsto (fun r => μ (sᶜ ∩ closedBall x r) / μ (closedBall x r)) (𝓝[>] 0) (𝓝 0) := by have A : Tendsto (fun r => μ (closedBall x r) / μ (closedBall x r)) (𝓝[>] 0) (𝓝 1) := by apply tendsto_const_nhds.congr' _ diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 033d11b1cb4a5..b61c2e17ffc68 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -214,7 +214,7 @@ theorem measure_add_diff (hs : NullMeasurableSet s μ) (t : Set α) : theorem measure_diff' (s : Set α) (hm : NullMeasurableSet t μ) (h_fin : μ t ≠ ∞) : μ (s \ t) = μ (s ∪ t) - μ t := - Eq.symm <| ENNReal.sub_eq_of_add_eq h_fin <| by rw [add_comm, measure_add_diff hm, union_comm] + ENNReal.eq_sub_of_add_eq h_fin <| by rw [add_comm, measure_add_diff hm, union_comm] theorem measure_diff (h : s₂ ⊆ s₁) (h₂ : NullMeasurableSet s₂ μ) (h_fin : μ s₂ ≠ ∞) : μ (s₁ \ s₂) = μ s₁ - μ s₂ := by rw [measure_diff' _ h₂ h_fin, union_eq_self_of_subset_right h] From ac1e7e8390055d66e3d62c8805e9fdefe79f5565 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 1 Oct 2024 21:09:42 +0000 Subject: [PATCH 096/174] =?UTF-8?q?feat:=20Composition=20of=20the=20coerci?= =?UTF-8?q?on=20`=E2=84=9D=20=E2=86=92=20=E2=84=82`=20with=20algebraic=20o?= =?UTF-8?q?perations=20(#17213)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanAPAP --- Mathlib/Algebra/Group/Hom/Defs.lean | 2 +- Mathlib/Data/Complex/Basic.lean | 22 ++++++++++++++++++++++ 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index f495bba039807..febcba6ed781c 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -452,7 +452,7 @@ theorem map_pow [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (a : G) : | n + 1 => by rw [pow_succ, pow_succ, map_mul, map_pow f a n] @[to_additive (attr := simp)] -lemma map_comp_pow [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : ι → G) (n : ℕ) : +lemma map_comp_pow [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (g : ι → G) (n : ℕ) : f ∘ (g ^ n) = f ∘ g ^ n := by ext; simp @[to_additive] diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 089de7d56ab00..bc2017ba047a1 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -637,6 +637,28 @@ def ofReal : ℝ →+* ℂ where theorem ofReal_eq_coe (r : ℝ) : ofReal r = r := rfl +variable {α : Type*} + +@[simp] lemma ofReal_comp_add (f g : α → ℝ) : ofReal' ∘ (f + g) = ofReal' ∘ f + ofReal' ∘ g := + map_comp_add ofReal .. + +@[simp] lemma ofReal_comp_sub (f g : α → ℝ) : ofReal' ∘ (f - g) = ofReal' ∘ f - ofReal' ∘ g := + map_comp_sub ofReal .. + +@[simp] lemma ofReal_comp_neg (f : α → ℝ) : ofReal' ∘ (-f) = -(ofReal' ∘ f) := map_comp_neg ofReal _ + +lemma ofReal_comp_nsmul (n : ℕ) (f : α → ℝ) : ofReal' ∘ (n • f) = n • (ofReal' ∘ f) := + map_comp_nsmul ofReal .. + +lemma ofReal_comp_zsmul (n : ℤ) (f : α → ℝ) : ofReal' ∘ (n • f) = n • (ofReal' ∘ f) := + map_comp_zsmul ofReal .. + +@[simp] lemma ofReal_comp_mul (f g : α → ℝ) : ofReal' ∘ (f * g) = ofReal' ∘ f * ofReal' ∘ g := + map_comp_mul ofReal .. + +@[simp] lemma ofReal_comp_pow (f : α → ℝ) (n : ℕ) : ofReal' ∘ (f ^ n) = (ofReal' ∘ f) ^ n := + map_comp_pow ofReal .. + @[simp] theorem I_sq : I ^ 2 = -1 := by rw [sq, I_mul_I] From a95d25fa5992f22ac4abe4219be331746411fca8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 1 Oct 2024 21:09:43 +0000 Subject: [PATCH 097/174] =?UTF-8?q?feat:=20`M=E2=81=BB=C2=B9.PosDef=20?= =?UTF-8?q?=E2=86=94=20M.PosDef`=20(#17310)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The implication already existed for `PosSemidef` Co-authored-by: Eric Wieser --- Mathlib/LinearAlgebra/Matrix/PosDef.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 941af97a89e99..9343bc8070623 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -450,6 +450,25 @@ theorem det_pos [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : 0 < det intro i _ simpa using hM.eigenvalues_pos i +theorem isUnit [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : IsUnit M := + isUnit_iff_isUnit_det _ |>.2 <| hM.det_pos.ne'.isUnit + +protected theorem inv [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : M⁻¹.PosDef := by + refine ⟨hM.isHermitian.inv, fun x hx => ?_⟩ + have := hM.2 (M⁻¹ *ᵥ x) ((Matrix.mulVec_injective_iff_isUnit.mpr ?_ |>.ne_iff' ?_).2 hx) + · let _inst := hM.isUnit.invertible + rwa [star_mulVec, mulVec_mulVec, Matrix.mul_inv_of_invertible, one_mulVec, + ← star_pos_iff, ← star_mulVec, ← star_dotProduct] at this + · simpa using hM.isUnit + · simp + +@[simp] +theorem _root_.Matrix.posDef_inv_iff [DecidableEq n] {M : Matrix n n 𝕜} : + M⁻¹.PosDef ↔ M.PosDef := + ⟨fun h => + letI := (Matrix.isUnit_nonsing_inv_iff.1 <| h.isUnit).invertible + Matrix.inv_inv_of_invertible M ▸ h.inv, (·.inv)⟩ + end PosDef end Matrix From 2ce1fb12cd6d7adcffb3b353244a8306bc026545 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 1 Oct 2024 21:09:44 +0000 Subject: [PATCH 098/174] chore: add a few focusing dots 2 (#17314) More missing cdots found by the multiGoal linter (#12339). --- Mathlib/Control/LawfulFix.lean | 3 +-- Mathlib/Data/Finsupp/Weight.lean | 2 +- Mathlib/Data/Seq/Computation.lean | 2 +- Mathlib/MeasureTheory/Function/UnifTight.lean | 3 ++- .../Measure/SeparableMeasure.lean | 10 +++++----- Mathlib/Order/Defs.lean | 18 ++++++++++++------ .../RingTheory/HahnSeries/Multiplication.lean | 8 ++++---- 7 files changed, 26 insertions(+), 20 deletions(-) diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index 485cc3ade329c..f22207cffe5c1 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -167,10 +167,9 @@ theorem fix_eq_ωSup_of_ωScottContinuous (hc : ωScottContinuous g) : Part.fix theorem fix_eq_of_ωScottContinuous (hc : ωScottContinuous g) : Part.fix g = g (Part.fix g) := by - rw [fix_eq_ωSup_of_ωScottContinuous, hc.map_ωSup] + rw [fix_eq_ωSup_of_ωScottContinuous hc, hc.map_ωSup] apply le_antisymm · apply ωSup_le_ωSup_of_le _ - exact hc intro i exists i intro x diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index fc57c22790c2d..ed767ee615ac8 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -157,7 +157,7 @@ theorem weight_eq_zero_iff_eq_zero ext s simp only [Finsupp.coe_zero, Pi.zero_apply] by_contra hs - apply NonTorsionWeight.ne_zero w _ + apply NonTorsionWeight.ne_zero w s rw [← nonpos_iff_eq_zero, ← h] exact le_weight_of_ne_zero' w hs · intro h diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index 2f8667ca00005..4622e5d162c96 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -695,7 +695,7 @@ theorem length_bind (s : Computation α) (f : α → Computation β) [_T1 : Term theorem of_results_bind {s : Computation α} {f : α → Computation β} {b k} : Results (bind s f) b k → ∃ a m n, Results s a m ∧ Results (f a) b n ∧ k = n + m := by induction k generalizing s with | zero => _ | succ n IH => _ - all_goals apply recOn s (fun a => _) fun s' => _ <;> intro e h + <;> apply recOn s (fun a => _) fun s' => _ <;> intro e h · simp only [ret_bind] at h exact ⟨e, _, _, results_pure _, h, rfl⟩ · have := congr_arg head (eq_thinkN h) diff --git a/Mathlib/MeasureTheory/Function/UnifTight.lean b/Mathlib/MeasureTheory/Function/UnifTight.lean index cc50e8c1de965..deaea6b197602 100644 --- a/Mathlib/MeasureTheory/Function/UnifTight.lean +++ b/Mathlib/MeasureTheory/Function/UnifTight.lean @@ -76,7 +76,8 @@ namespace UnifTight theorem eventually_cofinite_indicator (hf : UnifTight f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∀ᶠ s in μ.cofinite.smallSets, ∀ i, eLpNorm (s.indicator (f i)) p μ ≤ ε := by - by_cases hε_top : ε = ∞; subst hε_top; simp + by_cases hε_top : ε = ∞ + · subst hε_top; simp rcases hf (pos_iff_ne_zero.2 (toNNReal_ne_zero.mpr ⟨hε,hε_top⟩)) with ⟨s, hμs, hfs⟩ refine (eventually_smallSets' ?_).2 ⟨sᶜ, ?_, fun i ↦ (coe_toNNReal hε_top) ▸ hfs i⟩ · intro s t hst ht i diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index 55e3abb732208..ec687e378d87f 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -374,7 +374,6 @@ instance [CountablyGenerated X] [SFinite μ] : IsSeparable μ where ne_top_of_le_ne_top hμs <| μ.restrict_le_self _ rcases h𝒜.approx s ms this ε ε_pos with ⟨t, t_mem, ht⟩ refine ⟨t ∩ μ.sigmaFiniteSet, ⟨t, t_mem, rfl⟩, ?_⟩ - rw [← measure_inter_add_diff _ measurableSet_sigmaFiniteSet] have : μ (s ∆ (t ∩ μ.sigmaFiniteSet) \ μ.sigmaFiniteSet) = 0 := by rw [diff_eq_compl_inter, inter_symmDiff_distrib_left, ← ENNReal.bot_eq_zero, eq_bot_iff] calc @@ -384,10 +383,11 @@ instance [CountablyGenerated X] [SFinite μ] : IsSeparable μ where _ ≤ μ (μ.sigmaFiniteSetᶜ ∩ s) + μ (μ.sigmaFiniteSetᶜ ∩ (t ∩ μ.sigmaFiniteSet)) := measure_union_le _ _ _ = 0 := by - rw [inter_comm, ← μ.restrict_apply ms, hs, ← inter_assoc, inter_comm, ← inter_assoc, - inter_compl_self, empty_inter, measure_empty, zero_add] - rwa [this, add_zero, inter_symmDiff_distrib_right, inter_assoc, inter_self, - ← inter_symmDiff_distrib_right, ← μ.restrict_apply' measurableSet_sigmaFiniteSet] + rw [inter_comm, ← μ.restrict_apply ms, hs, ← inter_assoc, inter_comm, + ← inter_assoc, inter_compl_self, empty_inter, measure_empty, zero_add] + rwa [← measure_inter_add_diff _ measurableSet_sigmaFiniteSet, this, add_zero, + inter_symmDiff_distrib_right, inter_assoc, inter_self, ← inter_symmDiff_distrib_right, + ← μ.restrict_apply' measurableSet_sigmaFiniteSet] · refine False.elim <| hμs ?_ rw [eq_top_iff, ← hs] exact μ.restrict_le_self _ diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 93e79ca56a566..d50b65de7d6d3 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -502,9 +502,12 @@ lemma min_comm (a b : α) : min a b = min b a := lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by apply eq_min · apply le_trans; apply min_le_left; apply min_le_left - · apply le_min; apply le_trans; apply min_le_left; apply min_le_right; apply min_le_right - · intro d h₁ h₂; apply le_min; apply le_min h₁; apply le_trans h₂; apply min_le_left - apply le_trans h₂; apply min_le_right + · apply le_min + · apply le_trans; apply min_le_left; apply min_le_right + · apply min_le_right + · intro d h₁ h₂; apply le_min + · apply le_min h₁; apply le_trans h₂; apply min_le_left + · apply le_trans h₂; apply min_le_right lemma min_left_comm (a b c : α) : min a (min b c) = min b (min a c) := by rw [← min_assoc, min_comm a, min_assoc] @@ -526,9 +529,12 @@ lemma max_comm (a b : α) : max a b = max b a := lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by apply eq_max · apply le_trans; apply le_max_left a b; apply le_max_left - · apply max_le; apply le_trans; apply le_max_right a b; apply le_max_left; apply le_max_right - · intro d h₁ h₂; apply max_le; apply max_le h₁; apply le_trans (le_max_left _ _) h₂ - apply le_trans (le_max_right _ _) h₂ + · apply max_le + · apply le_trans; apply le_max_right a b; apply le_max_left + · apply le_max_right + · intro d h₁ h₂; apply max_le + · apply max_le h₁; apply le_trans (le_max_left _ _) h₂ + · apply le_trans (le_max_right _ _) h₂ lemma max_left_comm (a b c : α) : max a (max b c) = max b (max a c) := by rw [← max_assoc, max_comm a, max_assoc] diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index f3c1ef6a86bc5..153ac24993c5d 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -229,10 +229,10 @@ theorem add_smul [AddCommMonoid R] [SMulWithZero R V] {x y : HahnSeries Γ R} ext a have hwf := x.isPWO_support.union y.isPWO_support rw [smul_coeff_left hwf, HahnSeries.add_coeff', of_symm_add] - simp_all only [Pi.add_apply, HahnSeries.add_coeff'] - rw [smul_coeff_left hwf Set.subset_union_right, - smul_coeff_left hwf Set.subset_union_left] - · simp only [HahnSeries.add_coeff, h, sum_add_distrib] + · simp_all only [Pi.add_apply, HahnSeries.add_coeff'] + rw [smul_coeff_left hwf Set.subset_union_right, + smul_coeff_left hwf Set.subset_union_left] + simp only [HahnSeries.add_coeff, h, sum_add_distrib] · intro b simp_all only [Set.isPWO_union, HahnSeries.isPWO_support, and_self, HahnSeries.mem_support, HahnSeries.add_coeff, ne_eq, Set.mem_union, Set.mem_setOf_eq, mem_support] From 20048df56db5026323a309b7d7c3458d39aec6ce Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 1 Oct 2024 21:09:45 +0000 Subject: [PATCH 099/174] chore: remove repeated variables (#17339) --- Mathlib/Init/Algebra/Classes.lean | 4 ++-- Mathlib/Logic/Relator.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Init/Algebra/Classes.lean b/Mathlib/Init/Algebra/Classes.lean index 98fbccbde3cf3..c7be944592e5c 100644 --- a/Mathlib/Init/Algebra/Classes.lean +++ b/Mathlib/Init/Algebra/Classes.lean @@ -57,7 +57,7 @@ instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeak section -variable {α : Sort u} {r : α → α → Prop} +variable {r : α → α → Prop} local infixl:50 " ≺ " => r @@ -83,7 +83,7 @@ namespace StrictWeakOrder section -variable {α : Sort u} {r : α → α → Prop} +variable {r : α → α → Prop} local infixl:50 " ≺ " => r diff --git a/Mathlib/Logic/Relator.lean b/Mathlib/Logic/Relator.lean index 83b0b7e2046b8..17ab1660850b2 100644 --- a/Mathlib/Logic/Relator.lean +++ b/Mathlib/Logic/Relator.lean @@ -116,7 +116,7 @@ lemma rel_eq {r : α → β → Prop} (hr : BiUnique r) : (r ⇒ r ⇒ (·↔·) open Function -variable {α : Type*} {r₁₁ : α → α → Prop} {r₁₂ : α → β → Prop} {r₂₁ : β → α → Prop} +variable {r₁₁ : α → α → Prop} {r₁₂ : α → β → Prop} {r₂₁ : β → α → Prop} {r₂₃ : β → γ → Prop} {r₁₃ : α → γ → Prop} namespace LeftTotal From 9e4a4beff6f9d3ddf2c19c20948533afb4521927 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Tue, 1 Oct 2024 22:14:14 +0000 Subject: [PATCH 100/174] feat: a family of compatible triangularisable endomorphisms are simultaneously triangularisable (#17194) Co-authored-by: Oliver Nash <7734364+ocfnash@users.noreply.github.com> --- Mathlib/Algebra/Lie/Weights/Basic.lean | 35 ++----------- Mathlib/Algebra/Module/Submodule/Lattice.lean | 4 ++ .../Algebra/Module/Submodule/LinearMap.lean | 7 +++ Mathlib/Algebra/Module/Submodule/Map.lean | 4 ++ Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 45 +++++++++++++++++ .../Eigenspace/Triangularizable.lean | 49 +++++++++++++++++++ 6 files changed, 113 insertions(+), 31 deletions(-) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 310b0d062c7aa..810246b2c95a8 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -778,37 +778,10 @@ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizabl See also `LieModule.iSup_genWeightSpace_eq_top'`. -/ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by - generalize h_dim : finrank K M = n - induction n using Nat.strongRecOn generalizing M with | ind n ih => ?_ - obtain h' | ⟨y : L, hy : ¬ ∃ φ, genWeightSpaceOf M φ y = ⊤⟩ := - forall_or_exists_not (fun (x : L) ↦ ∃ (φ : K), genWeightSpaceOf M φ x = ⊤) - · choose χ hχ using h' - replace hχ : genWeightSpace M χ = ⊤ := by simpa only [genWeightSpace, hχ] using iInf_top - exact eq_top_iff.mpr <| hχ ▸ le_iSup (genWeightSpace M) χ - · replace hy : ∀ φ, finrank K (genWeightSpaceOf M φ y) < n := fun φ ↦ by - simp_rw [not_exists, ← lt_top_iff_ne_top] at hy; exact h_dim ▸ Submodule.finrank_lt (hy φ) - replace ih : ∀ φ, ⨆ χ : L → K, genWeightSpace (genWeightSpaceOf M φ y) χ = ⊤ := - fun φ ↦ ih _ (hy φ) (genWeightSpaceOf M φ y) rfl - replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), - genWeightSpace (genWeightSpaceOf M φ y) χ = ⊤ := by - intro φ - suffices ∀ χ : L → K, χ y ≠ φ → genWeightSpace (genWeightSpaceOf M φ y) χ = ⊥ by - specialize ih φ; rw [iSup_split, biSup_congr this] at ih; simpa using ih - intro χ hχ - rw [eq_bot_iff, ← (genWeightSpaceOf M φ y).ker_incl, LieModuleHom.ker, - ← LieSubmodule.map_le_iff_le_comap, map_genWeightSpace_eq_of_injective - (genWeightSpaceOf M φ y).injective_incl, LieSubmodule.range_incl, ← disjoint_iff_inf_le] - exact (disjoint_genWeightSpaceOf K L M hχ).mono_left - (genWeightSpace_le_genWeightSpaceOf M y χ) - replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), genWeightSpace M χ = genWeightSpaceOf M φ y := by - intro φ - have (χ : L → K) (hχ : χ y = φ) : genWeightSpace M χ = - (genWeightSpace (genWeightSpaceOf M φ y) χ).map (genWeightSpaceOf M φ y).incl := by - rw [← hχ, genWeightSpace_genWeightSpaceOf_map_incl] - simp_rw [biSup_congr this, ← LieSubmodule.map_iSup, ih, LieModuleHom.map_top, - LieSubmodule.range_incl] - simpa only [← ih, iSup_comm (ι := K), iSup_iSup_eq_right] using - iSup_genWeightSpaceOf_eq_top K L M y + simp only [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, + LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] + exact Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) + (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) (IsTriangularizable.iSup_eq_top) lemma iSup_genWeightSpace_eq_top' [IsTriangularizable K L M] : ⨆ χ : Weight K L M, genWeightSpace M χ = ⊤ := by diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index 1b9c97d71bbe6..ce617ef5b2a6b 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -241,6 +241,10 @@ theorem mem_finset_inf {ι} {s : Finset ι} {p : ι → Submodule R M} {x : M} : x ∈ s.inf p ↔ ∀ i ∈ s, x ∈ p i := by simp only [← SetLike.mem_coe, finset_inf_coe, Set.mem_iInter] +lemma inf_iInf {ι : Type*} [Nonempty ι] {p : ι → Submodule R M} (q : Submodule R M) : + q ⊓ ⨅ i, p i = ⨅ i, q ⊓ p i := + SetLike.coe_injective <| by simpa only [inf_coe, iInf_coe] using Set.inter_iInter _ _ + theorem mem_sup_left {S T : Submodule R M} : ∀ {x : M}, x ∈ S → x ∈ S ⊔ T := by have : S ≤ S ⊔ T := le_sup_left rw [LE.le] at this diff --git a/Mathlib/Algebra/Module/Submodule/LinearMap.lean b/Mathlib/Algebra/Module/Submodule/LinearMap.lean index 913de1199bb0c..36b04e8dc9c3d 100644 --- a/Mathlib/Algebra/Module/Submodule/LinearMap.lean +++ b/Mathlib/Algebra/Module/Submodule/LinearMap.lean @@ -183,6 +183,13 @@ lemma restrict_comp (g ∘ₗ f).restrict hfg = (g.restrict hg) ∘ₗ (f.restrict hf) := rfl +-- TODO Consider defining `Algebra R (p.compatibleMaps p)`, `AlgHom` version of `LinearMap.restrict` +lemma restrict_smul_one + {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] {p : Submodule R M} + (μ : R) (h : ∀ x ∈ p, (μ • (1 : Module.End R M)) x ∈ p := fun _ ↦ p.smul_mem μ) : + (μ • 1 : Module.End R M).restrict h = μ • (1 : Module.End R p) := + rfl + lemma restrict_commute {f g : M →ₗ[R] M} (h : Commute f g) {p : Submodule R M} (hf : MapsTo f p p) (hg : MapsTo g p p) : Commute (f.restrict hf) (g.restrict hg) := by diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 6f6227fd83734..1d1245ea336c6 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -120,6 +120,10 @@ theorem map_inf (f : F) {p q : Submodule R M} (hf : Injective f) : (p ⊓ q).map f = p.map f ⊓ q.map f := SetLike.coe_injective <| Set.image_inter hf +lemma map_iInf {ι : Type*} [Nonempty ι] {p : ι → Submodule R M} (f : F) (hf : Injective f) : + (⨅ i, p i).map f = ⨅ i, (p i).map f := + SetLike.coe_injective <| by simpa only [map_coe, iInf_coe] using hf.injOn.image_iInter_eq + theorem range_map_nonempty (N : Submodule R M) : (Set.range (fun ϕ => Submodule.map ϕ N : (M →ₛₗ[σ₁₂] M₂) → Submodule R₂ M₂)).Nonempty := ⟨_, Set.mem_range.mpr ⟨0, rfl⟩⟩ diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index e5c7fcb4b592f..1495ca7c3ce48 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -449,6 +449,51 @@ 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, ((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 _)] + rfl + +lemma mapsTo_restrict_maxGenEigenspace_restrict_of_mapsTo + {p : Submodule R M} (f g : End R M) (hf : MapsTo f p p) (hg : MapsTo g p p) {μ₁ μ₂ : R} + (h : MapsTo f (g.maxGenEigenspace μ₁) (g.maxGenEigenspace μ₂)) : + MapsTo (f.restrict hf) + (maxGenEigenspace (g.restrict hg) μ₁) + (maxGenEigenspace (g.restrict hg) μ₂) := by + intro x hx + simp_rw [SetLike.mem_coe, mem_maxGenEigenspace, ← LinearMap.restrict_smul_one _, + LinearMap.restrict_sub _, LinearMap.pow_restrict _, LinearMap.restrict_apply, + Submodule.mk_eq_zero, ← mem_maxGenEigenspace] at hx ⊢ + exact h hx + /-- If `p` is an invariant submodule of an endomorphism `f`, then the `μ`-eigenspace of the restriction of `f` to `p` is a submodule of the `μ`-eigenspace of `f`. -/ theorem eigenspace_restrict_le_eigenspace (f : End R M) {p : Submodule R M} (hfp : ∀ x ∈ p, f x ∈ p) diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index e2cd61a713f7f..a6fb4ee72e490 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -222,3 +222,52 @@ 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 μ) + · 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] + 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 f1843eb954a21599b1b417d356952f426db55208 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Tue, 1 Oct 2024 22:56:04 +0000 Subject: [PATCH 101/174] chore(Data/Finset): split lattice file (#17049) The file Data/Finset/Lattice contained two disjoint developments, with disjoint imports, forming one natural cut location. We also split max and min into a separate file. --- .../AscendingDescendingSequences.lean | 1 + Mathlib.lean | 2 + Mathlib/Data/Finset/Lattice.lean | 759 +----------------- Mathlib/Data/Finset/Max.lean | 543 +++++++++++++ Mathlib/Data/Finset/PiInduction.lean | 1 + Mathlib/Data/Finset/Sigma.lean | 1 + Mathlib/Data/Finset/Sort.lean | 2 +- Mathlib/Data/Fintype/Lattice.lean | 2 +- Mathlib/Data/Set/Finite.lean | 2 +- Mathlib/Order/CompleteLattice/Finset.lean | 247 ++++++ Mathlib/Order/CountableDenseLinearOrder.lean | 2 +- 11 files changed, 806 insertions(+), 756 deletions(-) create mode 100644 Mathlib/Data/Finset/Max.lean create mode 100644 Mathlib/Order/CompleteLattice/Finset.lean diff --git a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean index f4dbb1df5644d..df05d9256c5a5 100644 --- a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean +++ b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.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 -/ +import Mathlib.Data.Finset.Max import Mathlib.Data.Fintype.Powerset /-! diff --git a/Mathlib.lean b/Mathlib.lean index 4b4ebd9d3c7a4..b41a5e09ea2e7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2174,6 +2174,7 @@ import Mathlib.Data.Finset.Grade import Mathlib.Data.Finset.Image import Mathlib.Data.Finset.Interval import Mathlib.Data.Finset.Lattice +import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.MulAntidiagonal import Mathlib.Data.Finset.NAry import Mathlib.Data.Finset.NatAntidiagonal @@ -3572,6 +3573,7 @@ import Mathlib.Order.CompactlyGenerated.Intervals import Mathlib.Order.Compare import Mathlib.Order.CompleteBooleanAlgebra import Mathlib.Order.CompleteLattice +import Mathlib.Order.CompleteLattice.Finset import Mathlib.Order.CompleteLatticeIntervals import Mathlib.Order.CompletePartialOrder import Mathlib.Order.CompleteSublattice diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 3efa17777c830..b41a111701bda 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -5,17 +5,22 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Data.Finset.Fold -import Mathlib.Data.Finset.Option import Mathlib.Data.Finset.Pi import Mathlib.Data.Finset.Prod import Mathlib.Data.Multiset.Lattice import Mathlib.Data.Set.Lattice import Mathlib.Order.Hom.Lattice -import Mathlib.Order.Minimal import Mathlib.Order.Nat /-! # Lattice operations on finsets + +This file is concerned with folding binary lattice operations over finsets. + +For the special case of maximum and minimum of a finset, see Max.lean. + +See also SetLattice.lean, which is instead concerned with how big lattice or set operations behave +when indexed by a finset. -/ assert_not_exists OrderedCommMonoid @@ -1152,529 +1157,6 @@ theorem exists_mem_eq_inf [OrderTop α] (s : Finset ι) (h : s.Nonempty) (f : ι end LinearOrder -/-! ### max and min of finite sets -/ - - -section MaxMin - -variable [LinearOrder α] - -/-- Let `s` be a finset in a linear order. Then `s.max` is the maximum of `s` if `s` is not empty, -and `⊥` otherwise. It belongs to `WithBot α`. If you want to get an element of `α`, see -`s.max'`. -/ -protected def max (s : Finset α) : WithBot α := - sup s (↑) - -theorem max_eq_sup_coe {s : Finset α} : s.max = s.sup (↑) := - rfl - -theorem max_eq_sup_withBot (s : Finset α) : s.max = sup s (↑) := - rfl - -@[simp] -theorem max_empty : (∅ : Finset α).max = ⊥ := - rfl - -@[simp] -theorem max_insert {a : α} {s : Finset α} : (insert a s).max = max ↑a s.max := - fold_insert_idem - -@[simp] -theorem max_singleton {a : α} : Finset.max {a} = (a : WithBot α) := by - rw [← insert_emptyc_eq] - exact max_insert - -theorem max_of_mem {s : Finset α} {a : α} (h : a ∈ s) : ∃ b : α, s.max = b := by - obtain ⟨b, h, _⟩ := le_sup (α := WithBot α) h _ rfl - exact ⟨b, h⟩ - -theorem max_of_nonempty {s : Finset α} (h : s.Nonempty) : ∃ a : α, s.max = a := - let ⟨_, h⟩ := h - max_of_mem h - -theorem max_eq_bot {s : Finset α} : s.max = ⊥ ↔ s = ∅ := - ⟨fun h ↦ s.eq_empty_or_nonempty.elim id fun H ↦ by - obtain ⟨a, ha⟩ := max_of_nonempty H - rw [h] at ha; cases ha; , -- the `;` is needed since the `cases` syntax allows `cases a, b` - fun h ↦ h.symm ▸ max_empty⟩ - -theorem mem_of_max {s : Finset α} : ∀ {a : α}, s.max = a → a ∈ s := by - induction' s using Finset.induction_on with b s _ ih - · intro _ H; cases H - · intro a h - by_cases p : b = a - · induction p - exact mem_insert_self b s - · cases' max_choice (↑b) s.max with q q <;> rw [max_insert, q] at h - · cases h - cases p rfl - · exact mem_insert_of_mem (ih h) - -theorem le_max {a : α} {s : Finset α} (as : a ∈ s) : ↑a ≤ s.max := - le_sup as - -theorem not_mem_of_max_lt_coe {a : α} {s : Finset α} (h : s.max < a) : a ∉ s := - mt le_max h.not_le - -theorem le_max_of_eq {s : Finset α} {a b : α} (h₁ : a ∈ s) (h₂ : s.max = b) : a ≤ b := - WithBot.coe_le_coe.mp <| (le_max h₁).trans h₂.le - -theorem not_mem_of_max_lt {s : Finset α} {a b : α} (h₁ : b < a) (h₂ : s.max = ↑b) : a ∉ s := - Finset.not_mem_of_max_lt_coe <| h₂.trans_lt <| WithBot.coe_lt_coe.mpr h₁ - -@[gcongr] -theorem max_mono {s t : Finset α} (st : s ⊆ t) : s.max ≤ t.max := - sup_mono st - -protected theorem max_le {M : WithBot α} {s : Finset α} (st : ∀ a ∈ s, (a : WithBot α) ≤ M) : - s.max ≤ M := - Finset.sup_le st - -@[simp] -protected lemma max_le_iff {m : WithBot α} {s : Finset α} : s.max ≤ m ↔ ∀ a ∈ s, a ≤ m := - Finset.sup_le_iff - -@[simp] -protected lemma max_eq_top [OrderTop α] {s : Finset α} : s.max = ⊤ ↔ ⊤ ∈ s := - Finset.sup_eq_top_iff.trans <| by simp - -/-- Let `s` be a finset in a linear order. Then `s.min` is the minimum of `s` if `s` is not empty, -and `⊤` otherwise. It belongs to `WithTop α`. If you want to get an element of `α`, see -`s.min'`. -/ -protected def min (s : Finset α) : WithTop α := - inf s (↑) - -theorem min_eq_inf_withTop (s : Finset α) : s.min = inf s (↑) := - rfl - -@[simp] -theorem min_empty : (∅ : Finset α).min = ⊤ := - rfl - -@[simp] -theorem min_insert {a : α} {s : Finset α} : (insert a s).min = min (↑a) s.min := - fold_insert_idem - -@[simp] -theorem min_singleton {a : α} : Finset.min {a} = (a : WithTop α) := by - rw [← insert_emptyc_eq] - exact min_insert - -theorem min_of_mem {s : Finset α} {a : α} (h : a ∈ s) : ∃ b : α, s.min = b := by - obtain ⟨b, h, _⟩ := inf_le (α := WithTop α) h _ rfl - exact ⟨b, h⟩ - -theorem min_of_nonempty {s : Finset α} (h : s.Nonempty) : ∃ a : α, s.min = a := - let ⟨_, h⟩ := h - min_of_mem h - -@[simp] -theorem min_eq_top {s : Finset α} : s.min = ⊤ ↔ s = ∅ := by - simp [Finset.min, eq_empty_iff_forall_not_mem] - -theorem mem_of_min {s : Finset α} : ∀ {a : α}, s.min = a → a ∈ s := - @mem_of_max αᵒᵈ _ s - -theorem min_le {a : α} {s : Finset α} (as : a ∈ s) : s.min ≤ a := - inf_le as - -theorem not_mem_of_coe_lt_min {a : α} {s : Finset α} (h : ↑a < s.min) : a ∉ s := - mt min_le h.not_le - -theorem min_le_of_eq {s : Finset α} {a b : α} (h₁ : b ∈ s) (h₂ : s.min = a) : a ≤ b := - WithTop.coe_le_coe.mp <| h₂.ge.trans (min_le h₁) - -theorem not_mem_of_lt_min {s : Finset α} {a b : α} (h₁ : a < b) (h₂ : s.min = ↑b) : a ∉ s := - Finset.not_mem_of_coe_lt_min <| (WithTop.coe_lt_coe.mpr h₁).trans_eq h₂.symm - -@[gcongr] -theorem min_mono {s t : Finset α} (st : s ⊆ t) : t.min ≤ s.min := - inf_mono st - -protected theorem le_min {m : WithTop α} {s : Finset α} (st : ∀ a : α, a ∈ s → m ≤ a) : m ≤ s.min := - Finset.le_inf st - -@[simp] -protected theorem le_min_iff {m : WithTop α} {s : Finset α} : m ≤ s.min ↔ ∀ a ∈ s, m ≤ a := - Finset.le_inf_iff - -@[simp] -protected theorem min_eq_bot [OrderBot α] {s : Finset α} : s.min = ⊥ ↔ ⊥ ∈ s := - Finset.max_eq_top (α := αᵒᵈ) - -/-- Given a nonempty finset `s` in a linear order `α`, then `s.min' H` is its minimum, as an -element of `α`, where `H` is a proof of nonemptiness. Without this assumption, use instead `s.min`, -taking values in `WithTop α`. -/ -def min' (s : Finset α) (H : s.Nonempty) : α := - inf' s H id - -/-- Given a nonempty finset `s` in a linear order `α`, then `s.max' H` is its maximum, as an -element of `α`, where `H` is a proof of nonemptiness. Without this assumption, use instead `s.max`, -taking values in `WithBot α`. -/ -def max' (s : Finset α) (H : s.Nonempty) : α := - sup' s H id - -variable (s : Finset α) (H : s.Nonempty) {x : α} - -theorem min'_mem : s.min' H ∈ s := - mem_of_min <| by simp only [Finset.min, min', id_eq, coe_inf']; rfl - -theorem min'_le (x) (H2 : x ∈ s) : s.min' ⟨x, H2⟩ ≤ x := - min_le_of_eq H2 (WithTop.coe_untop _ _).symm - -theorem le_min' (x) (H2 : ∀ y ∈ s, x ≤ y) : x ≤ s.min' H := - H2 _ <| min'_mem _ _ - -theorem isLeast_min' : IsLeast (↑s) (s.min' H) := - ⟨min'_mem _ _, min'_le _⟩ - -@[simp] -theorem le_min'_iff {x} : x ≤ s.min' H ↔ ∀ y ∈ s, x ≤ y := - le_isGLB_iff (isLeast_min' s H).isGLB - -/-- `{a}.min' _` is `a`. -/ -@[simp] -theorem min'_singleton (a : α) : ({a} : Finset α).min' (singleton_nonempty _) = a := by simp [min'] - -theorem max'_mem : s.max' H ∈ s := - mem_of_max <| by simp only [max', Finset.max, id_eq, coe_sup']; rfl - -theorem le_max' (x) (H2 : x ∈ s) : x ≤ s.max' ⟨x, H2⟩ := - le_max_of_eq H2 (WithBot.coe_unbot _ _).symm - -theorem max'_le (x) (H2 : ∀ y ∈ s, y ≤ x) : s.max' H ≤ x := - H2 _ <| max'_mem _ _ - -theorem isGreatest_max' : IsGreatest (↑s) (s.max' H) := - ⟨max'_mem _ _, le_max' _⟩ - -@[simp] -theorem max'_le_iff {x} : s.max' H ≤ x ↔ ∀ y ∈ s, y ≤ x := - isLUB_le_iff (isGreatest_max' s H).isLUB - -@[simp] -theorem max'_lt_iff {x} : s.max' H < x ↔ ∀ y ∈ s, y < x := - ⟨fun Hlt y hy => (s.le_max' y hy).trans_lt Hlt, fun H => H _ <| s.max'_mem _⟩ - -@[simp] -theorem lt_min'_iff : x < s.min' H ↔ ∀ y ∈ s, x < y := - @max'_lt_iff αᵒᵈ _ _ H _ - -theorem max'_eq_sup' : s.max' H = s.sup' H id := rfl - -theorem min'_eq_inf' : s.min' H = s.inf' H id := rfl - -/-- `{a}.max' _` is `a`. -/ -@[simp] -theorem max'_singleton (a : α) : ({a} : Finset α).max' (singleton_nonempty _) = a := by simp [max'] - -theorem min'_lt_max' {i j} (H1 : i ∈ s) (H2 : j ∈ s) (H3 : i ≠ j) : - s.min' ⟨i, H1⟩ < s.max' ⟨i, H1⟩ := - isGLB_lt_isLUB_of_ne (s.isLeast_min' _).isGLB (s.isGreatest_max' _).isLUB H1 H2 H3 - -/-- If there's more than 1 element, the min' is less than the max'. An alternate version of -`min'_lt_max'` which is sometimes more convenient. --/ -theorem min'_lt_max'_of_card (h₂ : 1 < card s) : - s.min' (Finset.card_pos.1 <| by omega) < s.max' (Finset.card_pos.1 <| by omega) := by - rcases one_lt_card.1 h₂ with ⟨a, ha, b, hb, hab⟩ - exact s.min'_lt_max' ha hb hab - -theorem map_ofDual_min (s : Finset αᵒᵈ) : s.min.map ofDual = (s.image ofDual).max := by - rw [max_eq_sup_withBot, sup_image] - exact congr_fun Option.map_id _ - -theorem map_ofDual_max (s : Finset αᵒᵈ) : s.max.map ofDual = (s.image ofDual).min := by - rw [min_eq_inf_withTop, inf_image] - exact congr_fun Option.map_id _ - -theorem map_toDual_min (s : Finset α) : s.min.map toDual = (s.image toDual).max := by - rw [max_eq_sup_withBot, sup_image] - exact congr_fun Option.map_id _ - -theorem map_toDual_max (s : Finset α) : s.max.map toDual = (s.image toDual).min := by - rw [min_eq_inf_withTop, inf_image] - exact congr_fun Option.map_id _ - --- Porting note: new proofs without `convert` for the next four theorems. - -theorem ofDual_min' {s : Finset αᵒᵈ} (hs : s.Nonempty) : - ofDual (min' s hs) = max' (s.image ofDual) (hs.image _) := by - rw [← WithBot.coe_eq_coe] - simp only [min'_eq_inf', id_eq, ofDual_inf', Function.comp_apply, coe_sup', max'_eq_sup', - sup_image] - rfl - -theorem ofDual_max' {s : Finset αᵒᵈ} (hs : s.Nonempty) : - ofDual (max' s hs) = min' (s.image ofDual) (hs.image _) := by - rw [← WithTop.coe_eq_coe] - simp only [max'_eq_sup', id_eq, ofDual_sup', Function.comp_apply, coe_inf', min'_eq_inf', - inf_image] - rfl - -theorem toDual_min' {s : Finset α} (hs : s.Nonempty) : - toDual (min' s hs) = max' (s.image toDual) (hs.image _) := by - rw [← WithBot.coe_eq_coe] - simp only [min'_eq_inf', id_eq, toDual_inf', Function.comp_apply, coe_sup', max'_eq_sup', - sup_image] - rfl - -theorem toDual_max' {s : Finset α} (hs : s.Nonempty) : - toDual (max' s hs) = min' (s.image toDual) (hs.image _) := by - rw [← WithTop.coe_eq_coe] - simp only [max'_eq_sup', id_eq, toDual_sup', Function.comp_apply, coe_inf', min'_eq_inf', - inf_image] - rfl - -theorem max'_subset {s t : Finset α} (H : s.Nonempty) (hst : s ⊆ t) : - s.max' H ≤ t.max' (H.mono hst) := - le_max' _ _ (hst (s.max'_mem H)) - -theorem min'_subset {s t : Finset α} (H : s.Nonempty) (hst : s ⊆ t) : - t.min' (H.mono hst) ≤ s.min' H := - min'_le _ _ (hst (s.min'_mem H)) - -theorem max'_insert (a : α) (s : Finset α) (H : s.Nonempty) : - (insert a s).max' (s.insert_nonempty a) = max (s.max' H) a := - (isGreatest_max' _ _).unique <| by - rw [coe_insert, max_comm] - exact (isGreatest_max' _ _).insert _ - -theorem min'_insert (a : α) (s : Finset α) (H : s.Nonempty) : - (insert a s).min' (s.insert_nonempty a) = min (s.min' H) a := - (isLeast_min' _ _).unique <| by - rw [coe_insert, min_comm] - exact (isLeast_min' _ _).insert _ - -theorem lt_max'_of_mem_erase_max' [DecidableEq α] {a : α} (ha : a ∈ s.erase (s.max' H)) : - a < s.max' H := - lt_of_le_of_ne (le_max' _ _ (mem_of_mem_erase ha)) <| ne_of_mem_of_not_mem ha <| not_mem_erase _ _ - -theorem min'_lt_of_mem_erase_min' [DecidableEq α] {a : α} (ha : a ∈ s.erase (s.min' H)) : - s.min' H < a := - @lt_max'_of_mem_erase_max' αᵒᵈ _ s H _ a ha - -/-- To rewrite from right to left, use `Monotone.map_finset_max'`. -/ -@[simp] -theorem max'_image [LinearOrder β] {f : α → β} (hf : Monotone f) (s : Finset α) - (h : (s.image f).Nonempty) : (s.image f).max' h = f (s.max' h.of_image) := by - simp only [max', sup'_image] - exact .symm <| comp_sup'_eq_sup'_comp _ _ fun _ _ ↦ hf.map_max - -/-- A version of `Finset.max'_image` with LHS and RHS reversed. -Also, this version assumes that `s` is nonempty, not its image. -/ -lemma _root_.Monotone.map_finset_max' [LinearOrder β] {f : α → β} (hf : Monotone f) {s : Finset α} - (h : s.Nonempty) : f (s.max' h) = (s.image f).max' (h.image f) := - .symm <| max'_image hf .. - -/-- To rewrite from right to left, use `Monotone.map_finset_min'`. -/ -@[simp] -theorem min'_image [LinearOrder β] {f : α → β} (hf : Monotone f) (s : Finset α) - (h : (s.image f).Nonempty) : (s.image f).min' h = f (s.min' h.of_image) := by - simp only [min', inf'_image] - exact .symm <| comp_inf'_eq_inf'_comp _ _ fun _ _ ↦ hf.map_min - -/-- A version of `Finset.min'_image` with LHS and RHS reversed. -Also, this version assumes that `s` is nonempty, not its image. -/ -lemma _root_.Monotone.map_finset_min' [LinearOrder β] {f : α → β} (hf : Monotone f) {s : Finset α} - (h : s.Nonempty) : f (s.min' h) = (s.image f).min' (h.image f) := - .symm <| min'_image hf .. - -theorem coe_max' {s : Finset α} (hs : s.Nonempty) : ↑(s.max' hs) = s.max := - coe_sup' hs id - -theorem coe_min' {s : Finset α} (hs : s.Nonempty) : ↑(s.min' hs) = s.min := - coe_inf' hs id - -theorem max_mem_image_coe {s : Finset α} (hs : s.Nonempty) : - s.max ∈ (s.image (↑) : Finset (WithBot α)) := - mem_image.2 ⟨max' s hs, max'_mem _ _, coe_max' hs⟩ - -theorem min_mem_image_coe {s : Finset α} (hs : s.Nonempty) : - s.min ∈ (s.image (↑) : Finset (WithTop α)) := - mem_image.2 ⟨min' s hs, min'_mem _ _, coe_min' hs⟩ - -theorem max_mem_insert_bot_image_coe (s : Finset α) : - s.max ∈ (insert ⊥ (s.image (↑)) : Finset (WithBot α)) := - mem_insert.2 <| s.eq_empty_or_nonempty.imp max_eq_bot.2 max_mem_image_coe - -theorem min_mem_insert_top_image_coe (s : Finset α) : - s.min ∈ (insert ⊤ (s.image (↑)) : Finset (WithTop α)) := - mem_insert.2 <| s.eq_empty_or_nonempty.imp min_eq_top.2 min_mem_image_coe - -theorem max'_erase_ne_self {s : Finset α} (s0 : (s.erase x).Nonempty) : (s.erase x).max' s0 ≠ x := - ne_of_mem_erase (max'_mem _ s0) - -theorem min'_erase_ne_self {s : Finset α} (s0 : (s.erase x).Nonempty) : (s.erase x).min' s0 ≠ x := - ne_of_mem_erase (min'_mem _ s0) - -theorem max_erase_ne_self {s : Finset α} : (s.erase x).max ≠ x := by - by_cases s0 : (s.erase x).Nonempty - · refine ne_of_eq_of_ne (coe_max' s0).symm ?_ - exact WithBot.coe_eq_coe.not.mpr (max'_erase_ne_self _) - · rw [not_nonempty_iff_eq_empty.mp s0, max_empty] - exact WithBot.bot_ne_coe - -theorem min_erase_ne_self {s : Finset α} : (s.erase x).min ≠ x := by - -- Porting note: old proof `convert @max_erase_ne_self αᵒᵈ _ _ _` - convert @max_erase_ne_self αᵒᵈ _ (toDual x) (s.map toDual.toEmbedding) using 1 - apply congr_arg -- Porting note: forces unfolding to see `Finset.min` is `Finset.max` - congr! - ext; simp only [mem_map_equiv]; exact Iff.rfl - -theorem exists_next_right {x : α} {s : Finset α} (h : ∃ y ∈ s, x < y) : - ∃ y ∈ s, x < y ∧ ∀ z ∈ s, x < z → y ≤ z := - have Hne : (s.filter (x < ·)).Nonempty := h.imp fun y hy => mem_filter.2 (by simpa) - have aux := mem_filter.1 (min'_mem _ Hne) - ⟨min' _ Hne, aux.1, by simp, fun z hzs hz => min'_le _ _ <| mem_filter.2 ⟨hzs, by simpa⟩⟩ - -theorem exists_next_left {x : α} {s : Finset α} (h : ∃ y ∈ s, y < x) : - ∃ y ∈ s, y < x ∧ ∀ z ∈ s, z < x → z ≤ y := - @exists_next_right αᵒᵈ _ x s h - -/-- If finsets `s` and `t` are interleaved, then `Finset.card s ≤ Finset.card t + 1`. -/ -theorem card_le_of_interleaved {s t : Finset α} - (h : ∀ᵉ (x ∈ s) (y ∈ s), - x < y → (∀ z ∈ s, z ∉ Set.Ioo x y) → ∃ z ∈ t, x < z ∧ z < y) : - s.card ≤ t.card + 1 := by - replace h : ∀ᵉ (x ∈ s) (y ∈ s), x < y → ∃ z ∈ t, x < z ∧ z < y := by - intro x hx y hy hxy - rcases exists_next_right ⟨y, hy, hxy⟩ with ⟨a, has, hxa, ha⟩ - rcases h x hx a has hxa fun z hzs hz => hz.2.not_le <| ha _ hzs hz.1 with ⟨b, hbt, hxb, hba⟩ - exact ⟨b, hbt, hxb, hba.trans_le <| ha _ hy hxy⟩ - set f : α → WithTop α := fun x => (t.filter fun y => x < y).min - have f_mono : StrictMonoOn f s := by - intro x hx y hy hxy - rcases h x hx y hy hxy with ⟨a, hat, hxa, hay⟩ - calc - f x ≤ a := min_le (mem_filter.2 ⟨hat, by simpa⟩) - _ < f y := - (Finset.lt_inf_iff <| WithTop.coe_lt_top a).2 fun b hb => - WithTop.coe_lt_coe.2 <| hay.trans (by simpa using (mem_filter.1 hb).2) - - calc - s.card = (s.image f).card := (card_image_of_injOn f_mono.injOn).symm - _ ≤ (insert ⊤ (t.image (↑)) : Finset (WithTop α)).card := - card_mono <| image_subset_iff.2 fun x _ => - insert_subset_insert _ (image_subset_image <| filter_subset _ _) - (min_mem_insert_top_image_coe _) - _ ≤ t.card + 1 := (card_insert_le _ _).trans (Nat.add_le_add_right card_image_le _) - -/-- If finsets `s` and `t` are interleaved, then `Finset.card s ≤ Finset.card (t \ s) + 1`. -/ -theorem card_le_diff_of_interleaved {s t : Finset α} - (h : - ∀ᵉ (x ∈ s) (y ∈ s), - x < y → (∀ z ∈ s, z ∉ Set.Ioo x y) → ∃ z ∈ t, x < z ∧ z < y) : - s.card ≤ (t \ s).card + 1 := - card_le_of_interleaved fun x hx y hy hxy hs => - let ⟨z, hzt, hxz, hzy⟩ := h x hx y hy hxy hs - ⟨z, mem_sdiff.2 ⟨hzt, fun hzs => hs z hzs ⟨hxz, hzy⟩⟩, hxz, hzy⟩ - -/-- Induction principle for `Finset`s in a linearly ordered type: a predicate is true on all -`s : Finset α` provided that: - -* it is true on the empty `Finset`, -* for every `s : Finset α` and an element `a` strictly greater than all elements of `s`, `p s` - implies `p (insert a s)`. -/ -@[elab_as_elim] -theorem induction_on_max [DecidableEq α] {p : Finset α → Prop} (s : Finset α) (h0 : p ∅) - (step : ∀ a s, (∀ x ∈ s, x < a) → p s → p (insert a s)) : p s := by - induction' s using Finset.strongInductionOn with s ihs - rcases s.eq_empty_or_nonempty with (rfl | hne) - · exact h0 - · have H : s.max' hne ∈ s := max'_mem s hne - rw [← insert_erase H] - exact step _ _ (fun x => s.lt_max'_of_mem_erase_max' hne) (ihs _ <| erase_ssubset H) - -/-- Induction principle for `Finset`s in a linearly ordered type: a predicate is true on all -`s : Finset α` provided that: - -* it is true on the empty `Finset`, -* for every `s : Finset α` and an element `a` strictly less than all elements of `s`, `p s` - implies `p (insert a s)`. -/ -@[elab_as_elim] -theorem induction_on_min [DecidableEq α] {p : Finset α → Prop} (s : Finset α) (h0 : p ∅) - (step : ∀ a s, (∀ x ∈ s, a < x) → p s → p (insert a s)) : p s := - @induction_on_max αᵒᵈ _ _ _ s h0 step - -end MaxMin - -section MaxMinInductionValue - -variable [LinearOrder α] [LinearOrder β] - -/-- Induction principle for `Finset`s in any type from which a given function `f` maps to a linearly -ordered type : a predicate is true on all `s : Finset α` provided that: - -* it is true on the empty `Finset`, -* for every `s : Finset α` and an element `a` such that for elements of `s` denoted by `x` we have - `f x ≤ f a`, `p s` implies `p (insert a s)`. -/ -@[elab_as_elim] -theorem induction_on_max_value [DecidableEq ι] (f : ι → α) {p : Finset ι → Prop} (s : Finset ι) - (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f x ≤ f a) → p s → p (insert a s)) : p s := by - induction' s using Finset.strongInductionOn with s ihs - rcases (s.image f).eq_empty_or_nonempty with (hne | hne) - · simp only [image_eq_empty] at hne - simp only [hne, h0] - · have H : (s.image f).max' hne ∈ s.image f := max'_mem (s.image f) hne - simp only [mem_image, exists_prop] at H - rcases H with ⟨a, has, hfa⟩ - rw [← insert_erase has] - refine step _ _ (not_mem_erase a s) (fun x hx => ?_) (ihs _ <| erase_ssubset has) - rw [hfa] - exact le_max' _ _ (mem_image_of_mem _ <| mem_of_mem_erase hx) - -/-- Induction principle for `Finset`s in any type from which a given function `f` maps to a linearly -ordered type : a predicate is true on all `s : Finset α` provided that: - -* it is true on the empty `Finset`, -* for every `s : Finset α` and an element `a` such that for elements of `s` denoted by `x` we have - `f a ≤ f x`, `p s` implies `p (insert a s)`. -/ -@[elab_as_elim] -theorem induction_on_min_value [DecidableEq ι] (f : ι → α) {p : Finset ι → Prop} (s : Finset ι) - (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f a ≤ f x) → p s → p (insert a s)) : p s := - @induction_on_max_value αᵒᵈ ι _ _ _ _ s h0 step - -end MaxMinInductionValue - -section ExistsMaxMin - -variable [LinearOrder α] - -theorem exists_max_image (s : Finset β) (f : β → α) (h : s.Nonempty) : - ∃ x ∈ s, ∀ x' ∈ s, f x' ≤ f x := by - cases' max_of_nonempty (h.image f) with y hy - rcases mem_image.mp (mem_of_max hy) with ⟨x, hx, rfl⟩ - exact ⟨x, hx, fun x' hx' => le_max_of_eq (mem_image_of_mem f hx') hy⟩ - -theorem exists_min_image (s : Finset β) (f : β → α) (h : s.Nonempty) : - ∃ x ∈ s, ∀ x' ∈ s, f x ≤ f x' := - @exists_max_image αᵒᵈ β _ s f h - -end ExistsMaxMin - -theorem isGLB_iff_isLeast [LinearOrder α] (i : α) (s : Finset α) (hs : s.Nonempty) : - IsGLB (s : Set α) i ↔ IsLeast (↑s) i := by - refine ⟨fun his => ?_, IsLeast.isGLB⟩ - suffices i = min' s hs by - rw [this] - exact isLeast_min' s hs - rw [IsGLB, IsGreatest, mem_lowerBounds, mem_upperBounds] at his - exact le_antisymm (his.1 (Finset.min' s hs) (Finset.min'_mem s hs)) (his.2 _ (Finset.min'_le s)) - -theorem isLUB_iff_isGreatest [LinearOrder α] (i : α) (s : Finset α) (hs : s.Nonempty) : - IsLUB (s : Set α) i ↔ IsGreatest (↑s) i := - @isGLB_iff_isLeast αᵒᵈ _ i s hs - -theorem isGLB_mem [LinearOrder α] {i : α} (s : Finset α) (his : IsGLB (s : Set α) i) - (hs : s.Nonempty) : i ∈ s := by - rw [← mem_coe] - exact ((isGLB_iff_isLeast i s hs).mp his).1 - -theorem isLUB_mem [LinearOrder α] {i : α} (s : Finset α) (his : IsLUB (s : Set α) i) - (hs : s.Nonempty) : i ∈ s := - @isGLB_mem αᵒᵈ _ i s his hs - end Finset namespace Multiset @@ -1723,230 +1205,3 @@ theorem sup_singleton' [DecidableEq α] (s : Finset α) : s.sup singleton = s := (s.sup_singleton'' _).trans image_id end Finset - -section Lattice - -variable {ι' : Sort*} [CompleteLattice α] - -/-- Supremum of `s i`, `i : ι`, is equal to the supremum over `t : Finset ι` of suprema -`⨆ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `iSup_eq_iSup_finset'` for a version -that works for `ι : Sort*`. -/ -theorem iSup_eq_iSup_finset (s : ι → α) : ⨆ i, s i = ⨆ t : Finset ι, ⨆ i ∈ t, s i := by - classical - refine le_antisymm ?_ ?_ - · exact iSup_le fun b => le_iSup_of_le {b} <| le_iSup_of_le b <| le_iSup_of_le (by simp) <| le_rfl - · exact iSup_le fun t => iSup_le fun b => iSup_le fun _ => le_iSup _ _ - -/-- Supremum of `s i`, `i : ι`, is equal to the supremum over `t : Finset ι` of suprema -`⨆ i ∈ t, s i`. This version works for `ι : Sort*`. See `iSup_eq_iSup_finset` for a version -that assumes `ι : Type*` but has no `PLift`s. -/ -theorem iSup_eq_iSup_finset' (s : ι' → α) : - ⨆ i, s i = ⨆ t : Finset (PLift ι'), ⨆ i ∈ t, s (PLift.down i) := by - rw [← iSup_eq_iSup_finset, ← Equiv.plift.surjective.iSup_comp]; rfl - -/-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : Finset ι` of infima -`⨅ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `iInf_eq_iInf_finset'` for a version -that works for `ι : Sort*`. -/ -theorem iInf_eq_iInf_finset (s : ι → α) : ⨅ i, s i = ⨅ (t : Finset ι) (i ∈ t), s i := - @iSup_eq_iSup_finset αᵒᵈ _ _ _ - -/-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : Finset ι` of infima -`⨅ i ∈ t, s i`. This version works for `ι : Sort*`. See `iInf_eq_iInf_finset` for a version -that assumes `ι : Type*` but has no `PLift`s. -/ -theorem iInf_eq_iInf_finset' (s : ι' → α) : - ⨅ i, s i = ⨅ t : Finset (PLift ι'), ⨅ i ∈ t, s (PLift.down i) := - @iSup_eq_iSup_finset' αᵒᵈ _ _ _ - -end Lattice - -namespace Set - -variable {ι' : Sort*} - -/-- Union of an indexed family of sets `s : ι → Set α` is equal to the union of the unions -of finite subfamilies. This version assumes `ι : Type*`. See also `iUnion_eq_iUnion_finset'` for -a version that works for `ι : Sort*`. -/ -theorem iUnion_eq_iUnion_finset (s : ι → Set α) : ⋃ i, s i = ⋃ t : Finset ι, ⋃ i ∈ t, s i := - iSup_eq_iSup_finset s - -/-- Union of an indexed family of sets `s : ι → Set α` is equal to the union of the unions -of finite subfamilies. This version works for `ι : Sort*`. See also `iUnion_eq_iUnion_finset` for -a version that assumes `ι : Type*` but avoids `PLift`s in the right hand side. -/ -theorem iUnion_eq_iUnion_finset' (s : ι' → Set α) : - ⋃ i, s i = ⋃ t : Finset (PLift ι'), ⋃ i ∈ t, s (PLift.down i) := - iSup_eq_iSup_finset' s - -/-- Intersection of an indexed family of sets `s : ι → Set α` is equal to the intersection of the -intersections of finite subfamilies. This version assumes `ι : Type*`. See also -`iInter_eq_iInter_finset'` for a version that works for `ι : Sort*`. -/ -theorem iInter_eq_iInter_finset (s : ι → Set α) : ⋂ i, s i = ⋂ t : Finset ι, ⋂ i ∈ t, s i := - iInf_eq_iInf_finset s - -/-- Intersection of an indexed family of sets `s : ι → Set α` is equal to the intersection of the -intersections of finite subfamilies. This version works for `ι : Sort*`. See also -`iInter_eq_iInter_finset` for a version that assumes `ι : Type*` but avoids `PLift`s in the right -hand side. -/ -theorem iInter_eq_iInter_finset' (s : ι' → Set α) : - ⋂ i, s i = ⋂ t : Finset (PLift ι'), ⋂ i ∈ t, s (PLift.down i) := - iInf_eq_iInf_finset' s - -end Set - -namespace Finset - -section minimal - -variable [DecidableEq α] {P : Finset α → Prop} {s : Finset α} - -theorem maximal_iff_forall_insert (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : - Maximal P s ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by - simp only [Maximal, and_congr_right_iff] - exact fun _ ↦ ⟨fun h x hxs hx ↦ hxs <| h hx (subset_insert _ _) (mem_insert_self x s), - fun h t ht hst x hxt ↦ by_contra fun hxs ↦ h x hxs (hP ht (insert_subset hxt hst))⟩ - -theorem minimal_iff_forall_diff_singleton (hP : ∀ ⦃s t⦄, P t → t ⊆ s → P s) : - Minimal P s ↔ P s ∧ ∀ x ∈ s, ¬ P (s.erase x) where - mp h := ⟨h.prop, fun x hxs hx ↦ by simpa using h.le_of_le hx (erase_subset _ _) hxs⟩ - mpr h := ⟨h.1, fun t ht hts x hxs ↦ by_contra fun hxt ↦ - h.2 x hxs <| hP ht (subset_erase.2 ⟨hts, hxt⟩)⟩ - -end minimal - -/-! ### Interaction with big lattice/set operations -/ - -section Lattice - -theorem iSup_coe [SupSet β] (f : α → β) (s : Finset α) : ⨆ x ∈ (↑s : Set α), f x = ⨆ x ∈ s, f x := - rfl - -theorem iInf_coe [InfSet β] (f : α → β) (s : Finset α) : ⨅ x ∈ (↑s : Set α), f x = ⨅ x ∈ s, f x := - rfl - -variable [CompleteLattice β] - -theorem iSup_singleton (a : α) (s : α → β) : ⨆ x ∈ ({a} : Finset α), s x = s a := by simp - -theorem iInf_singleton (a : α) (s : α → β) : ⨅ x ∈ ({a} : Finset α), s x = s a := by simp - -theorem iSup_option_toFinset (o : Option α) (f : α → β) : ⨆ x ∈ o.toFinset, f x = ⨆ x ∈ o, f x := by - simp - -theorem iInf_option_toFinset (o : Option α) (f : α → β) : ⨅ x ∈ o.toFinset, f x = ⨅ x ∈ o, f x := - @iSup_option_toFinset _ βᵒᵈ _ _ _ - -variable [DecidableEq α] - -theorem iSup_union {f : α → β} {s t : Finset α} : - ⨆ x ∈ s ∪ t, f x = (⨆ x ∈ s, f x) ⊔ ⨆ x ∈ t, f x := by simp [iSup_or, iSup_sup_eq] - -theorem iInf_union {f : α → β} {s t : Finset α} : - ⨅ x ∈ s ∪ t, f x = (⨅ x ∈ s, f x) ⊓ ⨅ x ∈ t, f x := - @iSup_union α βᵒᵈ _ _ _ _ _ - -theorem iSup_insert (a : α) (s : Finset α) (t : α → β) : - ⨆ x ∈ insert a s, t x = t a ⊔ ⨆ x ∈ s, t x := by - rw [insert_eq] - simp only [iSup_union, Finset.iSup_singleton] - -theorem iInf_insert (a : α) (s : Finset α) (t : α → β) : - ⨅ x ∈ insert a s, t x = t a ⊓ ⨅ x ∈ s, t x := - @iSup_insert α βᵒᵈ _ _ _ _ _ - -theorem iSup_finset_image {f : γ → α} {g : α → β} {s : Finset γ} : - ⨆ x ∈ s.image f, g x = ⨆ y ∈ s, g (f y) := by rw [← iSup_coe, coe_image, iSup_image, iSup_coe] - -theorem iInf_finset_image {f : γ → α} {g : α → β} {s : Finset γ} : - ⨅ x ∈ s.image f, g x = ⨅ y ∈ s, g (f y) := by rw [← iInf_coe, coe_image, iInf_image, iInf_coe] - -theorem iSup_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : - ⨆ i ∈ insert x t, Function.update f x s i = s ⊔ ⨆ i ∈ t, f i := by - simp only [Finset.iSup_insert, update_same] - rcongr (i hi); apply update_noteq; rintro rfl; exact hx hi - -theorem iInf_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : - ⨅ i ∈ insert x t, update f x s i = s ⊓ ⨅ i ∈ t, f i := - @iSup_insert_update α βᵒᵈ _ _ _ _ f _ hx - -theorem iSup_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → β) : - ⨆ y ∈ s.biUnion t, f y = ⨆ (x ∈ s) (y ∈ t x), f y := by simp [@iSup_comm _ α, iSup_and] - -theorem iInf_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → β) : - ⨅ y ∈ s.biUnion t, f y = ⨅ (x ∈ s) (y ∈ t x), f y := - @iSup_biUnion _ βᵒᵈ _ _ _ _ _ _ - -end Lattice - -theorem set_biUnion_coe (s : Finset α) (t : α → Set β) : ⋃ x ∈ (↑s : Set α), t x = ⋃ x ∈ s, t x := - rfl - -theorem set_biInter_coe (s : Finset α) (t : α → Set β) : ⋂ x ∈ (↑s : Set α), t x = ⋂ x ∈ s, t x := - rfl - -theorem set_biUnion_singleton (a : α) (s : α → Set β) : ⋃ x ∈ ({a} : Finset α), s x = s a := - iSup_singleton a s - -theorem set_biInter_singleton (a : α) (s : α → Set β) : ⋂ x ∈ ({a} : Finset α), s x = s a := - iInf_singleton a s - -@[simp] -theorem set_biUnion_preimage_singleton (f : α → β) (s : Finset β) : - ⋃ y ∈ s, f ⁻¹' {y} = f ⁻¹' s := - Set.biUnion_preimage_singleton f s - -theorem set_biUnion_option_toFinset (o : Option α) (f : α → Set β) : - ⋃ x ∈ o.toFinset, f x = ⋃ x ∈ o, f x := - iSup_option_toFinset o f - -theorem set_biInter_option_toFinset (o : Option α) (f : α → Set β) : - ⋂ x ∈ o.toFinset, f x = ⋂ x ∈ o, f x := - iInf_option_toFinset o f - -theorem subset_set_biUnion_of_mem {s : Finset α} {f : α → Set β} {x : α} (h : x ∈ s) : - f x ⊆ ⋃ y ∈ s, f y := - show f x ≤ ⨆ y ∈ s, f y from le_iSup_of_le x <| by simp only [h, iSup_pos, le_refl] - -variable [DecidableEq α] - -theorem set_biUnion_union (s t : Finset α) (u : α → Set β) : - ⋃ x ∈ s ∪ t, u x = (⋃ x ∈ s, u x) ∪ ⋃ x ∈ t, u x := - iSup_union - -theorem set_biInter_inter (s t : Finset α) (u : α → Set β) : - ⋂ x ∈ s ∪ t, u x = (⋂ x ∈ s, u x) ∩ ⋂ x ∈ t, u x := - iInf_union - -theorem set_biUnion_insert (a : α) (s : Finset α) (t : α → Set β) : - ⋃ x ∈ insert a s, t x = t a ∪ ⋃ x ∈ s, t x := - iSup_insert a s t - -theorem set_biInter_insert (a : α) (s : Finset α) (t : α → Set β) : - ⋂ x ∈ insert a s, t x = t a ∩ ⋂ x ∈ s, t x := - iInf_insert a s t - -theorem set_biUnion_finset_image {f : γ → α} {g : α → Set β} {s : Finset γ} : - ⋃ x ∈ s.image f, g x = ⋃ y ∈ s, g (f y) := - iSup_finset_image - -theorem set_biInter_finset_image {f : γ → α} {g : α → Set β} {s : Finset γ} : - ⋂ x ∈ s.image f, g x = ⋂ y ∈ s, g (f y) := - iInf_finset_image - -theorem set_biUnion_insert_update {x : α} {t : Finset α} (f : α → Set β) {s : Set β} (hx : x ∉ t) : - ⋃ i ∈ insert x t, @update _ _ _ f x s i = s ∪ ⋃ i ∈ t, f i := - iSup_insert_update f hx - -theorem set_biInter_insert_update {x : α} {t : Finset α} (f : α → Set β) {s : Set β} (hx : x ∉ t) : - ⋂ i ∈ insert x t, @update _ _ _ f x s i = s ∩ ⋂ i ∈ t, f i := - iInf_insert_update f hx - -theorem set_biUnion_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → Set β) : - ⋃ y ∈ s.biUnion t, f y = ⋃ (x ∈ s) (y ∈ t x), f y := - iSup_biUnion s t f - -theorem set_biInter_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → Set β) : - ⋂ y ∈ s.biUnion t, f y = ⋂ (x ∈ s) (y ∈ t x), f y := - iInf_biUnion s t f - -end Finset - -set_option linter.style.longFile 2100 diff --git a/Mathlib/Data/Finset/Max.lean b/Mathlib/Data/Finset/Max.lean new file mode 100644 index 0000000000000..a972df5dc3ddc --- /dev/null +++ b/Mathlib/Data/Finset/Max.lean @@ -0,0 +1,543 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Mathlib.Data.Finset.Lattice + +/-! +# Maximum and minimum of finite sets +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists MonoidWithZero + +open Function Multiset OrderDual + +variable {F α β γ ι κ : Type*} + +namespace Finset + +/-! ### max and min of finite sets -/ + +section MaxMin + +variable [LinearOrder α] + +/-- Let `s` be a finset in a linear order. Then `s.max` is the maximum of `s` if `s` is not empty, +and `⊥` otherwise. It belongs to `WithBot α`. If you want to get an element of `α`, see +`s.max'`. -/ +protected def max (s : Finset α) : WithBot α := + sup s (↑) + +theorem max_eq_sup_coe {s : Finset α} : s.max = s.sup (↑) := + rfl + +theorem max_eq_sup_withBot (s : Finset α) : s.max = sup s (↑) := + rfl + +@[simp] +theorem max_empty : (∅ : Finset α).max = ⊥ := + rfl + +@[simp] +theorem max_insert {a : α} {s : Finset α} : (insert a s).max = max ↑a s.max := + fold_insert_idem + +@[simp] +theorem max_singleton {a : α} : Finset.max {a} = (a : WithBot α) := by + rw [← insert_emptyc_eq] + exact max_insert + +theorem max_of_mem {s : Finset α} {a : α} (h : a ∈ s) : ∃ b : α, s.max = b := by + obtain ⟨b, h, _⟩ := le_sup (α := WithBot α) h _ rfl + exact ⟨b, h⟩ + +theorem max_of_nonempty {s : Finset α} (h : s.Nonempty) : ∃ a : α, s.max = a := + let ⟨_, h⟩ := h + max_of_mem h + +theorem max_eq_bot {s : Finset α} : s.max = ⊥ ↔ s = ∅ := + ⟨fun h ↦ s.eq_empty_or_nonempty.elim id fun H ↦ by + obtain ⟨a, ha⟩ := max_of_nonempty H + rw [h] at ha; cases ha; , -- the `;` is needed since the `cases` syntax allows `cases a, b` + fun h ↦ h.symm ▸ max_empty⟩ + +theorem mem_of_max {s : Finset α} : ∀ {a : α}, s.max = a → a ∈ s := by + induction' s using Finset.induction_on with b s _ ih + · intro _ H; cases H + · intro a h + by_cases p : b = a + · induction p + exact mem_insert_self b s + · cases' max_choice (↑b) s.max with q q <;> rw [max_insert, q] at h + · cases h + cases p rfl + · exact mem_insert_of_mem (ih h) + +theorem le_max {a : α} {s : Finset α} (as : a ∈ s) : ↑a ≤ s.max := + le_sup as + +theorem not_mem_of_max_lt_coe {a : α} {s : Finset α} (h : s.max < a) : a ∉ s := + mt le_max h.not_le + +theorem le_max_of_eq {s : Finset α} {a b : α} (h₁ : a ∈ s) (h₂ : s.max = b) : a ≤ b := + WithBot.coe_le_coe.mp <| (le_max h₁).trans h₂.le + +theorem not_mem_of_max_lt {s : Finset α} {a b : α} (h₁ : b < a) (h₂ : s.max = ↑b) : a ∉ s := + Finset.not_mem_of_max_lt_coe <| h₂.trans_lt <| WithBot.coe_lt_coe.mpr h₁ + +@[gcongr] +theorem max_mono {s t : Finset α} (st : s ⊆ t) : s.max ≤ t.max := + sup_mono st + +protected theorem max_le {M : WithBot α} {s : Finset α} (st : ∀ a ∈ s, (a : WithBot α) ≤ M) : + s.max ≤ M := + Finset.sup_le st + +@[simp] +protected lemma max_le_iff {m : WithBot α} {s : Finset α} : s.max ≤ m ↔ ∀ a ∈ s, a ≤ m := + Finset.sup_le_iff + +@[simp] +protected lemma max_eq_top [OrderTop α] {s : Finset α} : s.max = ⊤ ↔ ⊤ ∈ s := + Finset.sup_eq_top_iff.trans <| by simp + +/-- Let `s` be a finset in a linear order. Then `s.min` is the minimum of `s` if `s` is not empty, +and `⊤` otherwise. It belongs to `WithTop α`. If you want to get an element of `α`, see +`s.min'`. -/ +protected def min (s : Finset α) : WithTop α := + inf s (↑) + +theorem min_eq_inf_withTop (s : Finset α) : s.min = inf s (↑) := + rfl + +@[simp] +theorem min_empty : (∅ : Finset α).min = ⊤ := + rfl + +@[simp] +theorem min_insert {a : α} {s : Finset α} : (insert a s).min = min (↑a) s.min := + fold_insert_idem + +@[simp] +theorem min_singleton {a : α} : Finset.min {a} = (a : WithTop α) := by + rw [← insert_emptyc_eq] + exact min_insert + +theorem min_of_mem {s : Finset α} {a : α} (h : a ∈ s) : ∃ b : α, s.min = b := by + obtain ⟨b, h, _⟩ := inf_le (α := WithTop α) h _ rfl + exact ⟨b, h⟩ + +theorem min_of_nonempty {s : Finset α} (h : s.Nonempty) : ∃ a : α, s.min = a := + let ⟨_, h⟩ := h + min_of_mem h + +@[simp] +theorem min_eq_top {s : Finset α} : s.min = ⊤ ↔ s = ∅ := by + simp [Finset.min, eq_empty_iff_forall_not_mem] + +theorem mem_of_min {s : Finset α} : ∀ {a : α}, s.min = a → a ∈ s := + @mem_of_max αᵒᵈ _ s + +theorem min_le {a : α} {s : Finset α} (as : a ∈ s) : s.min ≤ a := + inf_le as + +theorem not_mem_of_coe_lt_min {a : α} {s : Finset α} (h : ↑a < s.min) : a ∉ s := + mt min_le h.not_le + +theorem min_le_of_eq {s : Finset α} {a b : α} (h₁ : b ∈ s) (h₂ : s.min = a) : a ≤ b := + WithTop.coe_le_coe.mp <| h₂.ge.trans (min_le h₁) + +theorem not_mem_of_lt_min {s : Finset α} {a b : α} (h₁ : a < b) (h₂ : s.min = ↑b) : a ∉ s := + Finset.not_mem_of_coe_lt_min <| (WithTop.coe_lt_coe.mpr h₁).trans_eq h₂.symm + +@[gcongr] +theorem min_mono {s t : Finset α} (st : s ⊆ t) : t.min ≤ s.min := + inf_mono st + +protected theorem le_min {m : WithTop α} {s : Finset α} (st : ∀ a : α, a ∈ s → m ≤ a) : m ≤ s.min := + Finset.le_inf st + +@[simp] +protected theorem le_min_iff {m : WithTop α} {s : Finset α} : m ≤ s.min ↔ ∀ a ∈ s, m ≤ a := + Finset.le_inf_iff + +@[simp] +protected theorem min_eq_bot [OrderBot α] {s : Finset α} : s.min = ⊥ ↔ ⊥ ∈ s := + Finset.max_eq_top (α := αᵒᵈ) + +/-- Given a nonempty finset `s` in a linear order `α`, then `s.min' H` is its minimum, as an +element of `α`, where `H` is a proof of nonemptiness. Without this assumption, use instead `s.min`, +taking values in `WithTop α`. -/ +def min' (s : Finset α) (H : s.Nonempty) : α := + inf' s H id + +/-- Given a nonempty finset `s` in a linear order `α`, then `s.max' H` is its maximum, as an +element of `α`, where `H` is a proof of nonemptiness. Without this assumption, use instead `s.max`, +taking values in `WithBot α`. -/ +def max' (s : Finset α) (H : s.Nonempty) : α := + sup' s H id + +variable (s : Finset α) (H : s.Nonempty) {x : α} + +theorem min'_mem : s.min' H ∈ s := + mem_of_min <| by simp only [Finset.min, min', id_eq, coe_inf']; rfl + +theorem min'_le (x) (H2 : x ∈ s) : s.min' ⟨x, H2⟩ ≤ x := + min_le_of_eq H2 (WithTop.coe_untop _ _).symm + +theorem le_min' (x) (H2 : ∀ y ∈ s, x ≤ y) : x ≤ s.min' H := + H2 _ <| min'_mem _ _ + +theorem isLeast_min' : IsLeast (↑s) (s.min' H) := + ⟨min'_mem _ _, min'_le _⟩ + +@[simp] +theorem le_min'_iff {x} : x ≤ s.min' H ↔ ∀ y ∈ s, x ≤ y := + le_isGLB_iff (isLeast_min' s H).isGLB + +/-- `{a}.min' _` is `a`. -/ +@[simp] +theorem min'_singleton (a : α) : ({a} : Finset α).min' (singleton_nonempty _) = a := by simp [min'] + +theorem max'_mem : s.max' H ∈ s := + mem_of_max <| by simp only [max', Finset.max, id_eq, coe_sup']; rfl + +theorem le_max' (x) (H2 : x ∈ s) : x ≤ s.max' ⟨x, H2⟩ := + le_max_of_eq H2 (WithBot.coe_unbot _ _).symm + +theorem max'_le (x) (H2 : ∀ y ∈ s, y ≤ x) : s.max' H ≤ x := + H2 _ <| max'_mem _ _ + +theorem isGreatest_max' : IsGreatest (↑s) (s.max' H) := + ⟨max'_mem _ _, le_max' _⟩ + +@[simp] +theorem max'_le_iff {x} : s.max' H ≤ x ↔ ∀ y ∈ s, y ≤ x := + isLUB_le_iff (isGreatest_max' s H).isLUB + +@[simp] +theorem max'_lt_iff {x} : s.max' H < x ↔ ∀ y ∈ s, y < x := + ⟨fun Hlt y hy => (s.le_max' y hy).trans_lt Hlt, fun H => H _ <| s.max'_mem _⟩ + +@[simp] +theorem lt_min'_iff : x < s.min' H ↔ ∀ y ∈ s, x < y := + @max'_lt_iff αᵒᵈ _ _ H _ + +theorem max'_eq_sup' : s.max' H = s.sup' H id := rfl + +theorem min'_eq_inf' : s.min' H = s.inf' H id := rfl + +/-- `{a}.max' _` is `a`. -/ +@[simp] +theorem max'_singleton (a : α) : ({a} : Finset α).max' (singleton_nonempty _) = a := by simp [max'] + +theorem min'_lt_max' {i j} (H1 : i ∈ s) (H2 : j ∈ s) (H3 : i ≠ j) : + s.min' ⟨i, H1⟩ < s.max' ⟨i, H1⟩ := + isGLB_lt_isLUB_of_ne (s.isLeast_min' _).isGLB (s.isGreatest_max' _).isLUB H1 H2 H3 + +/-- If there's more than 1 element, the min' is less than the max'. An alternate version of +`min'_lt_max'` which is sometimes more convenient. +-/ +theorem min'_lt_max'_of_card (h₂ : 1 < card s) : + s.min' (Finset.card_pos.1 <| by omega) < s.max' (Finset.card_pos.1 <| by omega) := by + rcases one_lt_card.1 h₂ with ⟨a, ha, b, hb, hab⟩ + exact s.min'_lt_max' ha hb hab + +theorem map_ofDual_min (s : Finset αᵒᵈ) : s.min.map ofDual = (s.image ofDual).max := by + rw [max_eq_sup_withBot, sup_image] + exact congr_fun Option.map_id _ + +theorem map_ofDual_max (s : Finset αᵒᵈ) : s.max.map ofDual = (s.image ofDual).min := by + rw [min_eq_inf_withTop, inf_image] + exact congr_fun Option.map_id _ + +theorem map_toDual_min (s : Finset α) : s.min.map toDual = (s.image toDual).max := by + rw [max_eq_sup_withBot, sup_image] + exact congr_fun Option.map_id _ + +theorem map_toDual_max (s : Finset α) : s.max.map toDual = (s.image toDual).min := by + rw [min_eq_inf_withTop, inf_image] + exact congr_fun Option.map_id _ + +-- Porting note: new proofs without `convert` for the next four theorems. + +theorem ofDual_min' {s : Finset αᵒᵈ} (hs : s.Nonempty) : + ofDual (min' s hs) = max' (s.image ofDual) (hs.image _) := by + rw [← WithBot.coe_eq_coe] + simp only [min'_eq_inf', id_eq, ofDual_inf', Function.comp_apply, coe_sup', max'_eq_sup', + sup_image] + rfl + +theorem ofDual_max' {s : Finset αᵒᵈ} (hs : s.Nonempty) : + ofDual (max' s hs) = min' (s.image ofDual) (hs.image _) := by + rw [← WithTop.coe_eq_coe] + simp only [max'_eq_sup', id_eq, ofDual_sup', Function.comp_apply, coe_inf', min'_eq_inf', + inf_image] + rfl + +theorem toDual_min' {s : Finset α} (hs : s.Nonempty) : + toDual (min' s hs) = max' (s.image toDual) (hs.image _) := by + rw [← WithBot.coe_eq_coe] + simp only [min'_eq_inf', id_eq, toDual_inf', Function.comp_apply, coe_sup', max'_eq_sup', + sup_image] + rfl + +theorem toDual_max' {s : Finset α} (hs : s.Nonempty) : + toDual (max' s hs) = min' (s.image toDual) (hs.image _) := by + rw [← WithTop.coe_eq_coe] + simp only [max'_eq_sup', id_eq, toDual_sup', Function.comp_apply, coe_inf', min'_eq_inf', + inf_image] + rfl + +theorem max'_subset {s t : Finset α} (H : s.Nonempty) (hst : s ⊆ t) : + s.max' H ≤ t.max' (H.mono hst) := + le_max' _ _ (hst (s.max'_mem H)) + +theorem min'_subset {s t : Finset α} (H : s.Nonempty) (hst : s ⊆ t) : + t.min' (H.mono hst) ≤ s.min' H := + min'_le _ _ (hst (s.min'_mem H)) + +theorem max'_insert (a : α) (s : Finset α) (H : s.Nonempty) : + (insert a s).max' (s.insert_nonempty a) = max (s.max' H) a := + (isGreatest_max' _ _).unique <| by + rw [coe_insert, max_comm] + exact (isGreatest_max' _ _).insert _ + +theorem min'_insert (a : α) (s : Finset α) (H : s.Nonempty) : + (insert a s).min' (s.insert_nonempty a) = min (s.min' H) a := + (isLeast_min' _ _).unique <| by + rw [coe_insert, min_comm] + exact (isLeast_min' _ _).insert _ + +theorem lt_max'_of_mem_erase_max' [DecidableEq α] {a : α} (ha : a ∈ s.erase (s.max' H)) : + a < s.max' H := + lt_of_le_of_ne (le_max' _ _ (mem_of_mem_erase ha)) <| ne_of_mem_of_not_mem ha <| not_mem_erase _ _ + +theorem min'_lt_of_mem_erase_min' [DecidableEq α] {a : α} (ha : a ∈ s.erase (s.min' H)) : + s.min' H < a := + @lt_max'_of_mem_erase_max' αᵒᵈ _ s H _ a ha + +/-- To rewrite from right to left, use `Monotone.map_finset_max'`. -/ +@[simp] +theorem max'_image [LinearOrder β] {f : α → β} (hf : Monotone f) (s : Finset α) + (h : (s.image f).Nonempty) : (s.image f).max' h = f (s.max' h.of_image) := by + simp only [max', sup'_image] + exact .symm <| comp_sup'_eq_sup'_comp _ _ fun _ _ ↦ hf.map_max + +/-- A version of `Finset.max'_image` with LHS and RHS reversed. +Also, this version assumes that `s` is nonempty, not its image. -/ +lemma _root_.Monotone.map_finset_max' [LinearOrder β] {f : α → β} (hf : Monotone f) {s : Finset α} + (h : s.Nonempty) : f (s.max' h) = (s.image f).max' (h.image f) := + .symm <| max'_image hf .. + +/-- To rewrite from right to left, use `Monotone.map_finset_min'`. -/ +@[simp] +theorem min'_image [LinearOrder β] {f : α → β} (hf : Monotone f) (s : Finset α) + (h : (s.image f).Nonempty) : (s.image f).min' h = f (s.min' h.of_image) := by + simp only [min', inf'_image] + exact .symm <| comp_inf'_eq_inf'_comp _ _ fun _ _ ↦ hf.map_min + +/-- A version of `Finset.min'_image` with LHS and RHS reversed. +Also, this version assumes that `s` is nonempty, not its image. -/ +lemma _root_.Monotone.map_finset_min' [LinearOrder β] {f : α → β} (hf : Monotone f) {s : Finset α} + (h : s.Nonempty) : f (s.min' h) = (s.image f).min' (h.image f) := + .symm <| min'_image hf .. + +theorem coe_max' {s : Finset α} (hs : s.Nonempty) : ↑(s.max' hs) = s.max := + coe_sup' hs id + +theorem coe_min' {s : Finset α} (hs : s.Nonempty) : ↑(s.min' hs) = s.min := + coe_inf' hs id + +theorem max_mem_image_coe {s : Finset α} (hs : s.Nonempty) : + s.max ∈ (s.image (↑) : Finset (WithBot α)) := + mem_image.2 ⟨max' s hs, max'_mem _ _, coe_max' hs⟩ + +theorem min_mem_image_coe {s : Finset α} (hs : s.Nonempty) : + s.min ∈ (s.image (↑) : Finset (WithTop α)) := + mem_image.2 ⟨min' s hs, min'_mem _ _, coe_min' hs⟩ + +theorem max_mem_insert_bot_image_coe (s : Finset α) : + s.max ∈ (insert ⊥ (s.image (↑)) : Finset (WithBot α)) := + mem_insert.2 <| s.eq_empty_or_nonempty.imp max_eq_bot.2 max_mem_image_coe + +theorem min_mem_insert_top_image_coe (s : Finset α) : + s.min ∈ (insert ⊤ (s.image (↑)) : Finset (WithTop α)) := + mem_insert.2 <| s.eq_empty_or_nonempty.imp min_eq_top.2 min_mem_image_coe + +theorem max'_erase_ne_self {s : Finset α} (s0 : (s.erase x).Nonempty) : (s.erase x).max' s0 ≠ x := + ne_of_mem_erase (max'_mem _ s0) + +theorem min'_erase_ne_self {s : Finset α} (s0 : (s.erase x).Nonempty) : (s.erase x).min' s0 ≠ x := + ne_of_mem_erase (min'_mem _ s0) + +theorem max_erase_ne_self {s : Finset α} : (s.erase x).max ≠ x := by + by_cases s0 : (s.erase x).Nonempty + · refine ne_of_eq_of_ne (coe_max' s0).symm ?_ + exact WithBot.coe_eq_coe.not.mpr (max'_erase_ne_self _) + · rw [not_nonempty_iff_eq_empty.mp s0, max_empty] + exact WithBot.bot_ne_coe + +theorem min_erase_ne_self {s : Finset α} : (s.erase x).min ≠ x := by + -- Porting note: old proof `convert @max_erase_ne_self αᵒᵈ _ _ _` + convert @max_erase_ne_self αᵒᵈ _ (toDual x) (s.map toDual.toEmbedding) using 1 + apply congr_arg -- Porting note: forces unfolding to see `Finset.min` is `Finset.max` + congr! + ext; simp only [mem_map_equiv]; exact Iff.rfl + +theorem exists_next_right {x : α} {s : Finset α} (h : ∃ y ∈ s, x < y) : + ∃ y ∈ s, x < y ∧ ∀ z ∈ s, x < z → y ≤ z := + have Hne : (s.filter (x < ·)).Nonempty := h.imp fun y hy => mem_filter.2 (by simpa) + have aux := mem_filter.1 (min'_mem _ Hne) + ⟨min' _ Hne, aux.1, by simp, fun z hzs hz => min'_le _ _ <| mem_filter.2 ⟨hzs, by simpa⟩⟩ + +theorem exists_next_left {x : α} {s : Finset α} (h : ∃ y ∈ s, y < x) : + ∃ y ∈ s, y < x ∧ ∀ z ∈ s, z < x → z ≤ y := + @exists_next_right αᵒᵈ _ x s h + +/-- If finsets `s` and `t` are interleaved, then `Finset.card s ≤ Finset.card t + 1`. -/ +theorem card_le_of_interleaved {s t : Finset α} + (h : ∀ᵉ (x ∈ s) (y ∈ s), + x < y → (∀ z ∈ s, z ∉ Set.Ioo x y) → ∃ z ∈ t, x < z ∧ z < y) : + s.card ≤ t.card + 1 := by + replace h : ∀ᵉ (x ∈ s) (y ∈ s), x < y → ∃ z ∈ t, x < z ∧ z < y := by + intro x hx y hy hxy + rcases exists_next_right ⟨y, hy, hxy⟩ with ⟨a, has, hxa, ha⟩ + rcases h x hx a has hxa fun z hzs hz => hz.2.not_le <| ha _ hzs hz.1 with ⟨b, hbt, hxb, hba⟩ + exact ⟨b, hbt, hxb, hba.trans_le <| ha _ hy hxy⟩ + set f : α → WithTop α := fun x => (t.filter fun y => x < y).min + have f_mono : StrictMonoOn f s := by + intro x hx y hy hxy + rcases h x hx y hy hxy with ⟨a, hat, hxa, hay⟩ + calc + f x ≤ a := min_le (mem_filter.2 ⟨hat, by simpa⟩) + _ < f y := + (Finset.lt_inf_iff <| WithTop.coe_lt_top a).2 fun b hb => + WithTop.coe_lt_coe.2 <| hay.trans (by simpa using (mem_filter.1 hb).2) + + calc + s.card = (s.image f).card := (card_image_of_injOn f_mono.injOn).symm + _ ≤ (insert ⊤ (t.image (↑)) : Finset (WithTop α)).card := + card_mono <| image_subset_iff.2 fun x _ => + insert_subset_insert _ (image_subset_image <| filter_subset _ _) + (min_mem_insert_top_image_coe _) + _ ≤ t.card + 1 := (card_insert_le _ _).trans (Nat.add_le_add_right card_image_le _) + +/-- If finsets `s` and `t` are interleaved, then `Finset.card s ≤ Finset.card (t \ s) + 1`. -/ +theorem card_le_diff_of_interleaved {s t : Finset α} + (h : + ∀ᵉ (x ∈ s) (y ∈ s), + x < y → (∀ z ∈ s, z ∉ Set.Ioo x y) → ∃ z ∈ t, x < z ∧ z < y) : + s.card ≤ (t \ s).card + 1 := + card_le_of_interleaved fun x hx y hy hxy hs => + let ⟨z, hzt, hxz, hzy⟩ := h x hx y hy hxy hs + ⟨z, mem_sdiff.2 ⟨hzt, fun hzs => hs z hzs ⟨hxz, hzy⟩⟩, hxz, hzy⟩ + +/-- Induction principle for `Finset`s in a linearly ordered type: a predicate is true on all +`s : Finset α` provided that: + +* it is true on the empty `Finset`, +* for every `s : Finset α` and an element `a` strictly greater than all elements of `s`, `p s` + implies `p (insert a s)`. -/ +@[elab_as_elim] +theorem induction_on_max [DecidableEq α] {p : Finset α → Prop} (s : Finset α) (h0 : p ∅) + (step : ∀ a s, (∀ x ∈ s, x < a) → p s → p (insert a s)) : p s := by + induction' s using Finset.strongInductionOn with s ihs + rcases s.eq_empty_or_nonempty with (rfl | hne) + · exact h0 + · have H : s.max' hne ∈ s := max'_mem s hne + rw [← insert_erase H] + exact step _ _ (fun x => s.lt_max'_of_mem_erase_max' hne) (ihs _ <| erase_ssubset H) + +/-- Induction principle for `Finset`s in a linearly ordered type: a predicate is true on all +`s : Finset α` provided that: + +* it is true on the empty `Finset`, +* for every `s : Finset α` and an element `a` strictly less than all elements of `s`, `p s` + implies `p (insert a s)`. -/ +@[elab_as_elim] +theorem induction_on_min [DecidableEq α] {p : Finset α → Prop} (s : Finset α) (h0 : p ∅) + (step : ∀ a s, (∀ x ∈ s, a < x) → p s → p (insert a s)) : p s := + @induction_on_max αᵒᵈ _ _ _ s h0 step + +end MaxMin + +section MaxMinInductionValue + +variable [LinearOrder α] [LinearOrder β] + +/-- Induction principle for `Finset`s in any type from which a given function `f` maps to a linearly +ordered type : a predicate is true on all `s : Finset α` provided that: + +* it is true on the empty `Finset`, +* for every `s : Finset α` and an element `a` such that for elements of `s` denoted by `x` we have + `f x ≤ f a`, `p s` implies `p (insert a s)`. -/ +@[elab_as_elim] +theorem induction_on_max_value [DecidableEq ι] (f : ι → α) {p : Finset ι → Prop} (s : Finset ι) + (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f x ≤ f a) → p s → p (insert a s)) : p s := by + induction' s using Finset.strongInductionOn with s ihs + rcases (s.image f).eq_empty_or_nonempty with (hne | hne) + · simp only [image_eq_empty] at hne + simp only [hne, h0] + · have H : (s.image f).max' hne ∈ s.image f := max'_mem (s.image f) hne + simp only [mem_image, exists_prop] at H + rcases H with ⟨a, has, hfa⟩ + rw [← insert_erase has] + refine step _ _ (not_mem_erase a s) (fun x hx => ?_) (ihs _ <| erase_ssubset has) + rw [hfa] + exact le_max' _ _ (mem_image_of_mem _ <| mem_of_mem_erase hx) + +/-- Induction principle for `Finset`s in any type from which a given function `f` maps to a linearly +ordered type : a predicate is true on all `s : Finset α` provided that: + +* it is true on the empty `Finset`, +* for every `s : Finset α` and an element `a` such that for elements of `s` denoted by `x` we have + `f a ≤ f x`, `p s` implies `p (insert a s)`. -/ +@[elab_as_elim] +theorem induction_on_min_value [DecidableEq ι] (f : ι → α) {p : Finset ι → Prop} (s : Finset ι) + (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f a ≤ f x) → p s → p (insert a s)) : p s := + @induction_on_max_value αᵒᵈ ι _ _ _ _ s h0 step + +end MaxMinInductionValue + +section ExistsMaxMin + +variable [LinearOrder α] + +theorem exists_max_image (s : Finset β) (f : β → α) (h : s.Nonempty) : + ∃ x ∈ s, ∀ x' ∈ s, f x' ≤ f x := by + cases' max_of_nonempty (h.image f) with y hy + rcases mem_image.mp (mem_of_max hy) with ⟨x, hx, rfl⟩ + exact ⟨x, hx, fun x' hx' => le_max_of_eq (mem_image_of_mem f hx') hy⟩ + +theorem exists_min_image (s : Finset β) (f : β → α) (h : s.Nonempty) : + ∃ x ∈ s, ∀ x' ∈ s, f x ≤ f x' := + @exists_max_image αᵒᵈ β _ s f h + +end ExistsMaxMin + +theorem isGLB_iff_isLeast [LinearOrder α] (i : α) (s : Finset α) (hs : s.Nonempty) : + IsGLB (s : Set α) i ↔ IsLeast (↑s) i := by + refine ⟨fun his => ?_, IsLeast.isGLB⟩ + suffices i = min' s hs by + rw [this] + exact isLeast_min' s hs + rw [IsGLB, IsGreatest, mem_lowerBounds, mem_upperBounds] at his + exact le_antisymm (his.1 (Finset.min' s hs) (Finset.min'_mem s hs)) (his.2 _ (Finset.min'_le s)) + +theorem isLUB_iff_isGreatest [LinearOrder α] (i : α) (s : Finset α) (hs : s.Nonempty) : + IsLUB (s : Set α) i ↔ IsGreatest (↑s) i := + @isGLB_iff_isLeast αᵒᵈ _ i s hs + +theorem isGLB_mem [LinearOrder α] {i : α} (s : Finset α) (his : IsGLB (s : Set α) i) + (hs : s.Nonempty) : i ∈ s := by + rw [← mem_coe] + exact ((isGLB_iff_isLeast i s hs).mp his).1 + +theorem isLUB_mem [LinearOrder α] {i : α} (s : Finset α) (his : IsLUB (s : Set α) i) + (hs : s.Nonempty) : i ∈ s := + @isGLB_mem αᵒᵈ _ i s his hs + +end Finset diff --git a/Mathlib/Data/Finset/PiInduction.lean b/Mathlib/Data/Finset/PiInduction.lean index 4eb4a769d4e6a..14801945c227c 100644 --- a/Mathlib/Data/Finset/PiInduction.lean +++ b/Mathlib/Data/Finset/PiInduction.lean @@ -3,6 +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 -/ +import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.Sigma import Mathlib.Data.Fintype.Card diff --git a/Mathlib/Data/Finset/Sigma.lean b/Mathlib/Data/Finset/Sigma.lean index e3f0b82144225..9be90160317cf 100644 --- a/Mathlib/Data/Finset/Sigma.lean +++ b/Mathlib/Data/Finset/Sigma.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Yaël Dillies, Bhavik Mehta -/ import Mathlib.Data.Finset.Lattice import Mathlib.Data.Set.Sigma +import Mathlib.Order.CompleteLattice.Finset /-! # Finite sets in a sigma type diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 8d825b85b9622..dd1aab18ec416 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro import Mathlib.Order.RelIso.Set import Mathlib.Data.Multiset.Sort import Mathlib.Data.List.NodupEquivFin -import Mathlib.Data.Finset.Lattice +import Mathlib.Data.Finset.Max import Mathlib.Data.Fintype.Card /-! diff --git a/Mathlib/Data/Fintype/Lattice.lean b/Mathlib/Data/Fintype/Lattice.lean index 27835fd7979ce..64865d35bb0c7 100644 --- a/Mathlib/Data/Fintype/Lattice.lean +++ b/Mathlib/Data/Fintype/Lattice.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Fintype.Card -import Mathlib.Data.Finset.Lattice +import Mathlib.Data.Finset.Max /-! # Lemmas relating fintypes and order/lattice structure. diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index 3807e40b76e3c..2cc22fdd69e8c 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -3,8 +3,8 @@ 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, Kyle Miller -/ -import Mathlib.Data.Finset.Basic import Mathlib.Data.Finite.Basic +import Mathlib.Data.Finset.Max import Mathlib.Data.Set.Functor import Mathlib.Data.Set.Lattice diff --git a/Mathlib/Order/CompleteLattice/Finset.lean b/Mathlib/Order/CompleteLattice/Finset.lean new file mode 100644 index 0000000000000..b96b173756c62 --- /dev/null +++ b/Mathlib/Order/CompleteLattice/Finset.lean @@ -0,0 +1,247 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Mathlib.Data.Finset.Option +import Mathlib.Order.Minimal + +/-! +# Lattice operations on finsets + +This file is concerned with how big lattice or set operations behave when indexed by a finset. + +See also Lattice.lean, which is concerned with folding binary lattice operations over a finset. +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists MonoidWithZero + +open Function Multiset OrderDual + +variable {F α β γ ι κ : Type*} + +section Lattice + +variable {ι' : Sort*} [CompleteLattice α] + +/-- Supremum of `s i`, `i : ι`, is equal to the supremum over `t : Finset ι` of suprema +`⨆ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `iSup_eq_iSup_finset'` for a version +that works for `ι : Sort*`. -/ +theorem iSup_eq_iSup_finset (s : ι → α) : ⨆ i, s i = ⨆ t : Finset ι, ⨆ i ∈ t, s i := by + classical + refine le_antisymm ?_ ?_ + · exact iSup_le fun b => le_iSup_of_le {b} <| le_iSup_of_le b <| le_iSup_of_le (by simp) <| le_rfl + · exact iSup_le fun t => iSup_le fun b => iSup_le fun _ => le_iSup _ _ + +/-- Supremum of `s i`, `i : ι`, is equal to the supremum over `t : Finset ι` of suprema +`⨆ i ∈ t, s i`. This version works for `ι : Sort*`. See `iSup_eq_iSup_finset` for a version +that assumes `ι : Type*` but has no `PLift`s. -/ +theorem iSup_eq_iSup_finset' (s : ι' → α) : + ⨆ i, s i = ⨆ t : Finset (PLift ι'), ⨆ i ∈ t, s (PLift.down i) := by + rw [← iSup_eq_iSup_finset, ← Equiv.plift.surjective.iSup_comp]; rfl + +/-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : Finset ι` of infima +`⨅ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `iInf_eq_iInf_finset'` for a version +that works for `ι : Sort*`. -/ +theorem iInf_eq_iInf_finset (s : ι → α) : ⨅ i, s i = ⨅ (t : Finset ι) (i ∈ t), s i := + @iSup_eq_iSup_finset αᵒᵈ _ _ _ + +/-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : Finset ι` of infima +`⨅ i ∈ t, s i`. This version works for `ι : Sort*`. See `iInf_eq_iInf_finset` for a version +that assumes `ι : Type*` but has no `PLift`s. -/ +theorem iInf_eq_iInf_finset' (s : ι' → α) : + ⨅ i, s i = ⨅ t : Finset (PLift ι'), ⨅ i ∈ t, s (PLift.down i) := + @iSup_eq_iSup_finset' αᵒᵈ _ _ _ + +end Lattice + +namespace Set + +variable {ι' : Sort*} + +/-- Union of an indexed family of sets `s : ι → Set α` is equal to the union of the unions +of finite subfamilies. This version assumes `ι : Type*`. See also `iUnion_eq_iUnion_finset'` for +a version that works for `ι : Sort*`. -/ +theorem iUnion_eq_iUnion_finset (s : ι → Set α) : ⋃ i, s i = ⋃ t : Finset ι, ⋃ i ∈ t, s i := + iSup_eq_iSup_finset s + +/-- Union of an indexed family of sets `s : ι → Set α` is equal to the union of the unions +of finite subfamilies. This version works for `ι : Sort*`. See also `iUnion_eq_iUnion_finset` for +a version that assumes `ι : Type*` but avoids `PLift`s in the right hand side. -/ +theorem iUnion_eq_iUnion_finset' (s : ι' → Set α) : + ⋃ i, s i = ⋃ t : Finset (PLift ι'), ⋃ i ∈ t, s (PLift.down i) := + iSup_eq_iSup_finset' s + +/-- Intersection of an indexed family of sets `s : ι → Set α` is equal to the intersection of the +intersections of finite subfamilies. This version assumes `ι : Type*`. See also +`iInter_eq_iInter_finset'` for a version that works for `ι : Sort*`. -/ +theorem iInter_eq_iInter_finset (s : ι → Set α) : ⋂ i, s i = ⋂ t : Finset ι, ⋂ i ∈ t, s i := + iInf_eq_iInf_finset s + +/-- Intersection of an indexed family of sets `s : ι → Set α` is equal to the intersection of the +intersections of finite subfamilies. This version works for `ι : Sort*`. See also +`iInter_eq_iInter_finset` for a version that assumes `ι : Type*` but avoids `PLift`s in the right +hand side. -/ +theorem iInter_eq_iInter_finset' (s : ι' → Set α) : + ⋂ i, s i = ⋂ t : Finset (PLift ι'), ⋂ i ∈ t, s (PLift.down i) := + iInf_eq_iInf_finset' s + +end Set + +namespace Finset + +section minimal + +variable [DecidableEq α] {P : Finset α → Prop} {s : Finset α} + +theorem maximal_iff_forall_insert (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : + Maximal P s ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by + simp only [Maximal, and_congr_right_iff] + exact fun _ ↦ ⟨fun h x hxs hx ↦ hxs <| h hx (subset_insert _ _) (mem_insert_self x s), + fun h t ht hst x hxt ↦ by_contra fun hxs ↦ h x hxs (hP ht (insert_subset hxt hst))⟩ + +theorem minimal_iff_forall_diff_singleton (hP : ∀ ⦃s t⦄, P t → t ⊆ s → P s) : + Minimal P s ↔ P s ∧ ∀ x ∈ s, ¬ P (s.erase x) where + mp h := ⟨h.prop, fun x hxs hx ↦ by simpa using h.le_of_le hx (erase_subset _ _) hxs⟩ + mpr h := ⟨h.1, fun t ht hts x hxs ↦ by_contra fun hxt ↦ + h.2 x hxs <| hP ht (subset_erase.2 ⟨hts, hxt⟩)⟩ + +end minimal + +/-! ### Interaction with big lattice/set operations -/ + +section Lattice + +theorem iSup_coe [SupSet β] (f : α → β) (s : Finset α) : ⨆ x ∈ (↑s : Set α), f x = ⨆ x ∈ s, f x := + rfl + +theorem iInf_coe [InfSet β] (f : α → β) (s : Finset α) : ⨅ x ∈ (↑s : Set α), f x = ⨅ x ∈ s, f x := + rfl + +variable [CompleteLattice β] + +theorem iSup_singleton (a : α) (s : α → β) : ⨆ x ∈ ({a} : Finset α), s x = s a := by simp + +theorem iInf_singleton (a : α) (s : α → β) : ⨅ x ∈ ({a} : Finset α), s x = s a := by simp + +theorem iSup_option_toFinset (o : Option α) (f : α → β) : ⨆ x ∈ o.toFinset, f x = ⨆ x ∈ o, f x := by + simp + +theorem iInf_option_toFinset (o : Option α) (f : α → β) : ⨅ x ∈ o.toFinset, f x = ⨅ x ∈ o, f x := + @iSup_option_toFinset _ βᵒᵈ _ _ _ + +variable [DecidableEq α] + +theorem iSup_union {f : α → β} {s t : Finset α} : + ⨆ x ∈ s ∪ t, f x = (⨆ x ∈ s, f x) ⊔ ⨆ x ∈ t, f x := by simp [iSup_or, iSup_sup_eq] + +theorem iInf_union {f : α → β} {s t : Finset α} : + ⨅ x ∈ s ∪ t, f x = (⨅ x ∈ s, f x) ⊓ ⨅ x ∈ t, f x := + @iSup_union α βᵒᵈ _ _ _ _ _ + +theorem iSup_insert (a : α) (s : Finset α) (t : α → β) : + ⨆ x ∈ insert a s, t x = t a ⊔ ⨆ x ∈ s, t x := by + rw [insert_eq] + simp only [iSup_union, Finset.iSup_singleton] + +theorem iInf_insert (a : α) (s : Finset α) (t : α → β) : + ⨅ x ∈ insert a s, t x = t a ⊓ ⨅ x ∈ s, t x := + @iSup_insert α βᵒᵈ _ _ _ _ _ + +theorem iSup_finset_image {f : γ → α} {g : α → β} {s : Finset γ} : + ⨆ x ∈ s.image f, g x = ⨆ y ∈ s, g (f y) := by rw [← iSup_coe, coe_image, iSup_image, iSup_coe] + +theorem iInf_finset_image {f : γ → α} {g : α → β} {s : Finset γ} : + ⨅ x ∈ s.image f, g x = ⨅ y ∈ s, g (f y) := by rw [← iInf_coe, coe_image, iInf_image, iInf_coe] + +theorem iSup_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : + ⨆ i ∈ insert x t, Function.update f x s i = s ⊔ ⨆ i ∈ t, f i := by + simp only [Finset.iSup_insert, update_same] + rcongr (i hi); apply update_noteq; rintro rfl; exact hx hi + +theorem iInf_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : + ⨅ i ∈ insert x t, update f x s i = s ⊓ ⨅ i ∈ t, f i := + @iSup_insert_update α βᵒᵈ _ _ _ _ f _ hx + +theorem iSup_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → β) : + ⨆ y ∈ s.biUnion t, f y = ⨆ (x ∈ s) (y ∈ t x), f y := by simp [@iSup_comm _ α, iSup_and] + +theorem iInf_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → β) : + ⨅ y ∈ s.biUnion t, f y = ⨅ (x ∈ s) (y ∈ t x), f y := + @iSup_biUnion _ βᵒᵈ _ _ _ _ _ _ + +end Lattice + +theorem set_biUnion_coe (s : Finset α) (t : α → Set β) : ⋃ x ∈ (↑s : Set α), t x = ⋃ x ∈ s, t x := + rfl + +theorem set_biInter_coe (s : Finset α) (t : α → Set β) : ⋂ x ∈ (↑s : Set α), t x = ⋂ x ∈ s, t x := + rfl + +theorem set_biUnion_singleton (a : α) (s : α → Set β) : ⋃ x ∈ ({a} : Finset α), s x = s a := + iSup_singleton a s + +theorem set_biInter_singleton (a : α) (s : α → Set β) : ⋂ x ∈ ({a} : Finset α), s x = s a := + iInf_singleton a s + +@[simp] +theorem set_biUnion_preimage_singleton (f : α → β) (s : Finset β) : + ⋃ y ∈ s, f ⁻¹' {y} = f ⁻¹' s := + Set.biUnion_preimage_singleton f s + +theorem set_biUnion_option_toFinset (o : Option α) (f : α → Set β) : + ⋃ x ∈ o.toFinset, f x = ⋃ x ∈ o, f x := + iSup_option_toFinset o f + +theorem set_biInter_option_toFinset (o : Option α) (f : α → Set β) : + ⋂ x ∈ o.toFinset, f x = ⋂ x ∈ o, f x := + iInf_option_toFinset o f + +theorem subset_set_biUnion_of_mem {s : Finset α} {f : α → Set β} {x : α} (h : x ∈ s) : + f x ⊆ ⋃ y ∈ s, f y := + show f x ≤ ⨆ y ∈ s, f y from le_iSup_of_le x <| by simp only [h, iSup_pos, le_refl] + +variable [DecidableEq α] + +theorem set_biUnion_union (s t : Finset α) (u : α → Set β) : + ⋃ x ∈ s ∪ t, u x = (⋃ x ∈ s, u x) ∪ ⋃ x ∈ t, u x := + iSup_union + +theorem set_biInter_inter (s t : Finset α) (u : α → Set β) : + ⋂ x ∈ s ∪ t, u x = (⋂ x ∈ s, u x) ∩ ⋂ x ∈ t, u x := + iInf_union + +theorem set_biUnion_insert (a : α) (s : Finset α) (t : α → Set β) : + ⋃ x ∈ insert a s, t x = t a ∪ ⋃ x ∈ s, t x := + iSup_insert a s t + +theorem set_biInter_insert (a : α) (s : Finset α) (t : α → Set β) : + ⋂ x ∈ insert a s, t x = t a ∩ ⋂ x ∈ s, t x := + iInf_insert a s t + +theorem set_biUnion_finset_image {f : γ → α} {g : α → Set β} {s : Finset γ} : + ⋃ x ∈ s.image f, g x = ⋃ y ∈ s, g (f y) := + iSup_finset_image + +theorem set_biInter_finset_image {f : γ → α} {g : α → Set β} {s : Finset γ} : + ⋂ x ∈ s.image f, g x = ⋂ y ∈ s, g (f y) := + iInf_finset_image + +theorem set_biUnion_insert_update {x : α} {t : Finset α} (f : α → Set β) {s : Set β} (hx : x ∉ t) : + ⋃ i ∈ insert x t, @update _ _ _ f x s i = s ∪ ⋃ i ∈ t, f i := + iSup_insert_update f hx + +theorem set_biInter_insert_update {x : α} {t : Finset α} (f : α → Set β) {s : Set β} (hx : x ∉ t) : + ⋂ i ∈ insert x t, @update _ _ _ f x s i = s ∩ ⋂ i ∈ t, f i := + iInf_insert_update f hx + +theorem set_biUnion_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → Set β) : + ⋃ y ∈ s.biUnion t, f y = ⋃ (x ∈ s) (y ∈ t x), f y := + iSup_biUnion s t f + +theorem set_biInter_biUnion (s : Finset γ) (t : γ → Finset α) (f : α → Set β) : + ⋂ y ∈ s.biUnion t, f y = ⋂ (x ∈ s) (y ∈ t x), f y := + iInf_biUnion s t f + +end Finset diff --git a/Mathlib/Order/CountableDenseLinearOrder.lean b/Mathlib/Order/CountableDenseLinearOrder.lean index 1527b05b689ea..1ab3a1faebbc2 100644 --- a/Mathlib/Order/CountableDenseLinearOrder.lean +++ b/Mathlib/Order/CountableDenseLinearOrder.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ import Mathlib.Order.Ideal -import Mathlib.Data.Finset.Lattice +import Mathlib.Data.Finset.Max /-! # The back and forth method and countable dense linear orders From bc742f8b287472eff5fe02f324671ebf13eb3418 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Tue, 1 Oct 2024 22:56:05 +0000 Subject: [PATCH 102/174] feat(Combinatorics.SetFamily.Shadow): shadow_singleton (#17340) Co-authored-by: Moritz Firsching --- Mathlib/Combinatorics/SetFamily/Shadow.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/SetFamily/Shadow.lean b/Mathlib/Combinatorics/SetFamily/Shadow.lean index bc4f6f690c062..a3b293b819c4c 100644 --- a/Mathlib/Combinatorics/SetFamily/Shadow.lean +++ b/Mathlib/Combinatorics/SetFamily/Shadow.lean @@ -75,7 +75,10 @@ theorem shadow_empty : ∂ (∅ : Finset (Finset α)) = ∅ := theorem shadow_singleton_empty : ∂ ({∅} : Finset (Finset α)) = ∅ := rfl ---TODO: Prove `∂ {{a}} = {∅}` quickly using `covers` and `GradeOrder` +@[simp] +theorem shadow_singleton (a : α) : ∂ {{a}} = {∅} := by + simp [shadow] + /-- The shadow is monotone. -/ @[mono] theorem shadow_monotone : Monotone (shadow : Finset (Finset α) → Finset (Finset α)) := fun _ _ => From ebabb3402ee09467a5836072a47b298486eacf4a Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 2 Oct 2024 01:16:02 +0000 Subject: [PATCH 103/174] chore: add a few focusing dots 3 (#17313) More missing cdots found by the multiGoal linter (#12339). --- Mathlib/RingTheory/LocalProperties/Basic.lean | 4 ++-- Mathlib/RingTheory/MvPowerSeries/LexOrder.lean | 4 ++-- Mathlib/RingTheory/PowerSeries/Basic.lean | 16 ++++++++-------- 3 files changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/RingTheory/LocalProperties/Basic.lean b/Mathlib/RingTheory/LocalProperties/Basic.lean index 0141196dfed77..a3ece3c9a9726 100644 --- a/Mathlib/RingTheory/LocalProperties/Basic.lean +++ b/Mathlib/RingTheory/LocalProperties/Basic.lean @@ -240,8 +240,8 @@ theorem RingHom.PropertyIsLocal.ofLocalizationSpan (hP : RingHom.PropertyIsLocal rintro ⟨_, r, hr, rfl⟩ rw [← IsLocalization.map_comp (M := Submonoid.powers r) (S := Localization.Away r) (T := Submonoid.powers (f r))] - apply hP.StableUnderCompositionWithLocalizationAway.right _ r - exact hs' ⟨r, hr⟩ + · apply hP.StableUnderCompositionWithLocalizationAway.right _ r + exact hs' ⟨r, hr⟩ lemma RingHom.OfLocalizationSpanTarget.ofIsLocalization (hP : RingHom.OfLocalizationSpanTarget P) (hP' : RingHom.RespectsIso P) diff --git a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean index f5472cc79ef0c..13ecad6b47dce 100644 --- a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean +++ b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean @@ -34,8 +34,8 @@ noncomputable def lexOrder (φ : MvPowerSeries σ R) : (WithTop (Lex (σ →₀ simp only [Set.image_nonempty, Function.support_nonempty_iff, ne_eq, h, not_false_eq_true] apply WithTop.some apply WellFounded.min _ (toLex '' φ.support) ne - exact Finsupp.instLTLex.lt - exact wellFounded_lt + · exact Finsupp.instLTLex.lt + · exact wellFounded_lt theorem lexOrder_def_of_ne_zero {φ : MvPowerSeries σ R} (hφ : φ ≠ 0) : ∃ (ne : Set.Nonempty (toLex '' φ.support)), diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 33cea1c3cd017..57422669ec4cd 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -603,14 +603,14 @@ lemma coeff_one_pow (n : ℕ) (φ : R⟦X⟧) : CharP.cast_eq_zero, zero_add, mul_one, not_true_eq_false] at h'' norm_num at h'' · rw [ih] - conv => lhs; arg 2; rw [mul_comm, ← mul_assoc] - move_mul [← (constantCoeff R) φ ^ (n' - 1)] - conv => enter [1, 2, 1, 1, 2]; rw [← pow_one (a := constantCoeff R φ)] - rw [← pow_add (a := constantCoeff R φ)] - conv => enter [1, 2, 1, 1]; rw [Nat.sub_add_cancel h'] - conv => enter [1, 2, 1]; rw [mul_comm] - rw [mul_assoc, ← one_add_mul, add_comm, mul_assoc] - conv => enter [1, 2]; rw [mul_comm] + · conv => lhs; arg 2; rw [mul_comm, ← mul_assoc] + move_mul [← (constantCoeff R) φ ^ (n' - 1)] + conv => enter [1, 2, 1, 1, 2]; rw [← pow_one (a := constantCoeff R φ)] + rw [← pow_add (a := constantCoeff R φ)] + conv => enter [1, 2, 1, 1]; rw [Nat.sub_add_cancel h'] + conv => enter [1, 2, 1]; rw [mul_comm] + rw [mul_assoc, ← one_add_mul, add_comm, mul_assoc] + conv => enter [1, 2]; rw [mul_comm] exact h' · decide From 1235f45ece3c1f049543e58769e29b31340f7269 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 2 Oct 2024 04:12:44 +0000 Subject: [PATCH 104/174] chore: update Mathlib dependencies 2024-10-02 (#17347) 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 9161d8bacca4b..1c5880467526f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c9e106b0541a3b1c4bf82017aca78a50cc3e5ca2", + "rev": "781beceb959c68b36d3d96205b3531e341879d2c", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From 081ccd5754d78a2d08f4ec4a737e5c7b54e76743 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Wed, 2 Oct 2024 04:55:20 +0000 Subject: [PATCH 105/174] feat(Analysis/Normed/{Group,Field}/Ultra): Nonarchimedean norms from ultrametrics (#14768) Co-authored-by: Yakov Pechersky Co-authored-by: Yakov Pechersky Co-authored-by: Yury G. Kudryashov --- Mathlib.lean | 2 + Mathlib/Analysis/Normed/Group/Ultra.lean | 195 ++++++++++++++++++ Mathlib/Analysis/Normed/Ring/Ultra.lean | 78 +++++++ Mathlib/Topology/MetricSpace/Ultra/Basic.lean | 12 ++ 4 files changed, 287 insertions(+) create mode 100644 Mathlib/Analysis/Normed/Group/Ultra.lean create mode 100644 Mathlib/Analysis/Normed/Ring/Ultra.lean diff --git a/Mathlib.lean b/Mathlib.lean index b41a5e09ea2e7..96d5225a538f8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1232,6 +1232,7 @@ import Mathlib.Analysis.Normed.Group.SemiNormedGrp.Kernels import Mathlib.Analysis.Normed.Group.Seminorm import Mathlib.Analysis.Normed.Group.Submodule import Mathlib.Analysis.Normed.Group.Tannery +import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Analysis.Normed.Group.ZeroAtInfty import Mathlib.Analysis.Normed.Lp.LpEquiv @@ -1261,6 +1262,7 @@ import Mathlib.Analysis.Normed.Order.UpperLower import Mathlib.Analysis.Normed.Ring.Seminorm import Mathlib.Analysis.Normed.Ring.SeminormFromBounded import Mathlib.Analysis.Normed.Ring.SeminormFromConst +import Mathlib.Analysis.Normed.Ring.Ultra import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Analysis.NormedSpace.BallAction import Mathlib.Analysis.NormedSpace.ConformalLinearMap diff --git a/Mathlib/Analysis/Normed/Group/Ultra.lean b/Mathlib/Analysis/Normed/Group/Ultra.lean new file mode 100644 index 0000000000000..040fcc1e08cdc --- /dev/null +++ b/Mathlib/Analysis/Normed/Group/Ultra.lean @@ -0,0 +1,195 @@ +/- +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.Group.Basic +import Mathlib.Topology.MetricSpace.Ultra.Basic + +/-! +# Ultrametric norms + +This file contains results on the behavior of norms in ultrametric groups. + +## Main results + +* `IsUltrametricDist.isUltrametricDist_of_isNonarchimedean_norm`: + a normed additive group has an ultrametric iff the norm is nonarchimedean + +## Implementation details + +Some results are proved first about `nnnorm : X → ℝ≥0` because the bottom element +in `NNReal` is 0, so easier to make statements about maxima of empty sets. + +## Tags + +ultrametric, nonarchimedean +-/ +open Metric NNReal + +namespace IsUltrametricDist + +section Group + +variable {S S' ι : Type*} [SeminormedGroup S] [SeminormedGroup S'] [IsUltrametricDist S] + +@[to_additive] +lemma norm_mul_le_max (x y : S) : + ‖x * y‖ ≤ max ‖x‖ ‖y‖ := by + simpa only [le_max_iff, dist_eq_norm_div, div_inv_eq_mul, div_one, one_mul] using + dist_triangle_max x 1 y⁻¹ + +@[to_additive] +lemma isUltrametricDist_of_forall_norm_mul_le_max_norm + (h : ∀ x y : S', ‖x * y‖ ≤ max ‖x‖ ‖y‖) : IsUltrametricDist S' where + dist_triangle_max x y z := by + simpa only [dist_eq_norm_div, le_max_iff, div_mul_div_cancel] using h (x / y) (y / z) + +lemma isUltrametricDist_of_isNonarchimedean_norm {S' : Type*} [SeminormedAddGroup S'] + (h : IsNonarchimedean (norm : S' → ℝ)) : IsUltrametricDist S' := + isUltrametricDist_of_forall_norm_add_le_max_norm h + +@[to_additive] +lemma nnnorm_mul_le_max (x y : S) : + ‖x * y‖₊ ≤ max ‖x‖₊ ‖y‖₊ := + norm_mul_le_max _ _ + +@[to_additive] +lemma isUltrametricDist_of_forall_nnnorm_mul_le_max_nnnorm + (h : ∀ x y : S', ‖x * y‖₊ ≤ max ‖x‖₊ ‖y‖₊) : IsUltrametricDist S' := + isUltrametricDist_of_forall_norm_mul_le_max_norm h + +lemma isUltrametricDist_of_isNonarchimedean_nnnorm {S' : Type*} [SeminormedAddGroup S'] + (h : IsNonarchimedean ((↑) ∘ (nnnorm : S' → ℝ≥0))) : IsUltrametricDist S' := + isUltrametricDist_of_forall_nnnorm_add_le_max_nnnorm h + +/-- All triangles are isosceles in an ultrametric normed group. -/ +@[to_additive "All triangles are isosceles in an ultrametric normed additive group."] +lemma norm_mul_eq_max_of_norm_ne_norm + {x y : S} (h : ‖x‖ ≠ ‖y‖) : ‖x * y‖ = max ‖x‖ ‖y‖ := by + rw [← div_inv_eq_mul, ← dist_eq_norm_div, dist_eq_max_of_dist_ne_dist _ 1 _ (by simp [h])] + simp only [dist_one_right, dist_one_left, norm_inv'] + +@[to_additive] +lemma norm_eq_of_mul_norm_lt_max {x y : S} (h : ‖x * y‖ < max ‖x‖ ‖y‖) : + ‖x‖ = ‖y‖ := + not_ne_iff.mp (h.ne ∘ norm_mul_eq_max_of_norm_ne_norm) + +/-- All triangles are isosceles in an ultrametric normed group. -/ +@[to_additive "All triangles are isosceles in an ultrametric normed additive group."] +lemma nnnorm_mul_eq_max_of_nnnorm_ne_nnnorm + {x y : S} (h : ‖x‖₊ ≠ ‖y‖₊) : ‖x * y‖₊ = max ‖x‖₊ ‖y‖₊ := by + simpa only [← NNReal.coe_inj, NNReal.coe_max] using + norm_mul_eq_max_of_norm_ne_norm (NNReal.coe_injective.ne h) + +@[to_additive] +lemma nnnorm_eq_of_mul_nnnorm_lt_max {x y : S} (h : ‖x * y‖₊ < max ‖x‖₊ ‖y‖₊) : + ‖x‖₊ = ‖y‖₊ := + not_ne_iff.mp (h.ne ∘ nnnorm_mul_eq_max_of_nnnorm_ne_nnnorm) + +/-- All triangles are isosceles in an ultrametric normed group. -/ +@[to_additive "All triangles are isosceles in an ultrametric normed additive group."] +lemma norm_div_eq_max_of_norm_div_ne_norm_div (x y z : S) (h : ‖x / y‖ ≠ ‖y / z‖) : + ‖x / z‖ = max ‖x / y‖ ‖y / z‖ := by + simpa only [div_mul_div_cancel] using norm_mul_eq_max_of_norm_ne_norm h + +/-- All triangles are isosceles in an ultrametric normed group. -/ +@[to_additive "All triangles are isosceles in an ultrametric normed additive group."] +lemma nnnorm_div_eq_max_of_nnnorm_div_ne_nnnorm_div (x y z : S) (h : ‖x / y‖₊ ≠ ‖y / z‖₊) : + ‖x / z‖₊ = max ‖x / y‖₊ ‖y / z‖₊ := by + simpa only [← NNReal.coe_inj, NNReal.coe_max] using + norm_div_eq_max_of_norm_div_ne_norm_div _ _ _ (NNReal.coe_injective.ne h) + +@[to_additive] +lemma nnnorm_pow_le (x : S) (n : ℕ) : + ‖x ^ n‖₊ ≤ ‖x‖₊ := by + induction n with + | zero => simp + | succ n hn => simpa [pow_add, hn] using nnnorm_mul_le_max (x ^ n) x + +@[to_additive] +lemma norm_pow_le (x : S) (n : ℕ) : + ‖x ^ n‖ ≤ ‖x‖ := + nnnorm_pow_le x n + +@[to_additive] +lemma nnnorm_zpow_le (x : S) (z : ℤ) : + ‖x ^ z‖₊ ≤ ‖x‖₊ := by + cases z <;> + simpa using nnnorm_pow_le _ _ + +@[to_additive] +lemma norm_zpow_le (x : S) (z : ℤ) : + ‖x ^ z‖ ≤ ‖x‖ := + nnnorm_zpow_le x z + +end Group + +section CommGroup + +variable {M ι : Type*} [SeminormedCommGroup M] [IsUltrametricDist M] + +/-- Nonarchimedean norm of a product is less than or equal the norm of any term in the product. +This version is phrased using `Finset.sup'` and `Finset.Nonempty` due to `Finset.sup` +operating over an `OrderBot`, which `ℝ` is not. +-/ +@[to_additive "Nonarchimedean norm of a sum is less than or equal the norm of any term in the sum. +This version is phrased using `Finset.sup'` and `Finset.Nonempty` due to `Finset.sup` +operating over an `OrderBot`, which `ℝ` is not. "] +lemma _root_.Finset.Nonempty.norm_prod_le_sup'_norm {s : Finset ι} (hs : s.Nonempty) (f : ι → M) : + ‖∏ i ∈ s, f i‖ ≤ s.sup' hs (‖f ·‖) := by + simp only [Finset.le_sup'_iff] + induction hs using Finset.Nonempty.cons_induction with + | singleton j => simp only [Finset.mem_singleton, Finset.prod_singleton, exists_eq_left, le_refl] + | cons j t hj _ IH => + simp only [Finset.prod_cons, Finset.mem_cons, exists_eq_or_imp] + refine (le_total ‖∏ i ∈ t, f i‖ ‖f j‖).imp ?_ ?_ <;> intro h + · exact (norm_mul_le_max _ _).trans (max_eq_left h).le + · exact ⟨_, IH.choose_spec.left, (norm_mul_le_max _ _).trans <| + ((max_eq_right h).le.trans IH.choose_spec.right)⟩ + +/-- Nonarchimedean norm of a product is less than or equal to the largest norm of a term in the +product. -/ +@[to_additive "Nonarchimedean norm of a sum is less than or equal to the largest norm of a term in +the sum."] +lemma _root_.Finset.nnnorm_prod_le_sup_nnnorm (s : Finset ι) (f : ι → M) : + ‖∏ i ∈ s, f i‖₊ ≤ s.sup (‖f ·‖₊) := by + rcases s.eq_empty_or_nonempty with rfl|hs + · simp only [Finset.prod_empty, nnnorm_one', Finset.sup_empty, bot_eq_zero', le_refl] + · simpa only [← Finset.sup'_eq_sup hs, Finset.le_sup'_iff, coe_le_coe, coe_nnnorm'] + using hs.norm_prod_le_sup'_norm f + +/-- +Generalised ultrametric triangle inequality for finite products in commutative groups with +an ultrametric norm. +-/ +@[to_additive "Generalised ultrametric triangle inequality for finite sums in additive commutative +groups with an ultrametric norm."] +lemma nnnorm_prod_le_of_forall_le {s : Finset ι} {f : ι → M} {C : ℝ≥0} + (hC : ∀ i ∈ s, ‖f i‖₊ ≤ C) : ‖∏ i ∈ s, f i‖₊ ≤ C := + (s.nnnorm_prod_le_sup_nnnorm f).trans <| Finset.sup_le hC + +/-- +Generalised ultrametric triangle inequality for nonempty finite products in commutative groups with +an ultrametric norm. +-/ +@[to_additive "Generalised ultrametric triangle inequality for nonempty finite sums in additive +commutative groups with an ultrametric norm."] +lemma norm_prod_le_of_forall_le_of_nonempty {s : Finset ι} (hs : s.Nonempty) {f : ι → M} {C : ℝ} + (hC : ∀ i ∈ s, ‖f i‖ ≤ C) : ‖∏ i ∈ s, f i‖ ≤ C := + (hs.norm_prod_le_sup'_norm f).trans (Finset.sup'_le hs _ hC) + +/-- +Generalised ultrametric triangle inequality for finite products in commutative groups with +an ultrametric norm. +-/ +@[to_additive "Generalised ultrametric triangle inequality for finite sums in additive commutative +groups with an ultrametric norm."] +lemma norm_prod_le_of_forall_le_of_nonneg {s : Finset ι} {f : ι → M} {C : ℝ} + (h_nonneg : 0 ≤ C) (hC : ∀ i ∈ s, ‖f i‖ ≤ C) : ‖∏ i ∈ s, f i‖ ≤ C := by + lift C to NNReal using h_nonneg + exact nnnorm_prod_le_of_forall_le hC + +end CommGroup + +end IsUltrametricDist diff --git a/Mathlib/Analysis/Normed/Ring/Ultra.lean b/Mathlib/Analysis/Normed/Ring/Ultra.lean new file mode 100644 index 0000000000000..1e8016ac72edb --- /dev/null +++ b/Mathlib/Analysis/Normed/Ring/Ultra.lean @@ -0,0 +1,78 @@ +/- +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 + +/-! +# Ultrametric norms on rings where the norm of one is one + +This file contains results on the behavior of norms in ultrametric normed rings. +The norm must send one to one. + +## Main results + +* `norm_intCast_le_one`: + the norm of the image of an integer in the ring is always less than or equal to one + +## Implementation details + +A `[NormedRing R]` only assumes a submultiplicative norm and does not have `[NormOneClass R]`. +The weakest ring-like structure that has a bundled norm such that `‖1‖ = 1` is +`[NormedDivisionRing K]`. +Since the statements below hold in any context, we can state them +in an unbundled fashion using `[NormOneClass R]`. +In fact one can actually prove all these lemmas only assuming +`{R : Type*} [SeminormedAddGroup R] [One R] [NormOneClass R] [IsUltrametricDist R]`. +But one has to give the typeclass machinery a little help in order to get it to recognise that there +is a coercion from `ℕ` or `ℤ` to `R`. +Instead, we use weakest pre-existing typeclass that implies both +`[SeminormedAddGroup R]` and `[AddGroupWithOne R]`, which is `[SeminormedRing R]`. + +## Tags + +ultrametric, nonarchimedean +-/ +open Metric NNReal + +namespace IsUltrametricDist + +section NormOneClass + +variable {R : Type*} [SeminormedRing R] [NormOneClass R] [IsUltrametricDist R] + +lemma norm_add_one_le_max_norm_one (x : R) : + ‖x + 1‖ ≤ max ‖x‖ 1 := by + simpa only [le_max_iff, norm_one] using norm_add_le_max x 1 + +lemma nnnorm_add_one_le_max_nnnorm_one (x : R) : + ‖x + 1‖₊ ≤ max ‖x‖₊ 1 := + norm_add_one_le_max_norm_one _ + +variable (R) +lemma nnnorm_natCast_le_one (n : ℕ) : + ‖(n : R)‖₊ ≤ 1 := by + induction n with + | zero => simp only [Nat.cast_zero, nnnorm_zero, zero_le] + | succ n hn => simpa only [Nat.cast_add, Nat.cast_one, hn, max_eq_right] using + nnnorm_add_one_le_max_nnnorm_one (n : R) + +lemma norm_natCast_le_one (n : ℕ) : + ‖(n : R)‖ ≤ 1 := + nnnorm_natCast_le_one R n + +lemma nnnorm_intCast_le_one (z : ℤ) : + ‖(z : R)‖₊ ≤ 1 := by + induction z <;> + simpa only [Int.ofNat_eq_coe, Int.cast_natCast, Int.cast_negSucc, Nat.cast_one, nnnorm_neg] + using nnnorm_natCast_le_one _ _ + +lemma norm_intCast_le_one (z : ℤ) : + ‖(z : R)‖ ≤ 1 := + nnnorm_intCast_le_one _ z + +end NormOneClass + +end IsUltrametricDist diff --git a/Mathlib/Topology/MetricSpace/Ultra/Basic.lean b/Mathlib/Topology/MetricSpace/Ultra/Basic.lean index 5b00adb8cfda9..ceccb0c003e73 100644 --- a/Mathlib/Topology/MetricSpace/Ultra/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Ultra/Basic.lean @@ -50,6 +50,18 @@ lemma dist_triangle_max : dist x z ≤ max (dist x y) (dist y z) := namespace IsUltrametricDist +/-- All triangles are isosceles in an ultrametric space. -/ +lemma dist_eq_max_of_dist_ne_dist (h : dist x y ≠ dist y z) : + dist x z = max (dist x y) (dist y z) := by + apply le_antisymm (dist_triangle_max x y z) + rcases h.lt_or_lt with h | h + · rw [max_eq_right h.le] + apply (le_max_iff.mp <| dist_triangle_max y x z).resolve_left + simpa only [not_le, dist_comm x y] using h + · rw [max_eq_left h.le, dist_comm x y, dist_comm x z] + apply (le_max_iff.mp <| dist_triangle_max y z x).resolve_left + simpa only [not_le, dist_comm x y] using h + instance subtype (p : X → Prop) : IsUltrametricDist (Subtype p) := ⟨fun _ _ _ ↦ by simpa [Subtype.dist_eq] using dist_triangle_max _ _ _⟩ From 9593a645885b383cc56d09543a34af0c44c695bc Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 2 Oct 2024 05:35:43 +0000 Subject: [PATCH 106/174] =?UTF-8?q?feat(NumberTheory/{GaussSum|JacobiSum/B?= =?UTF-8?q?asic}):=20formula=20for=20g(=CF=87)^ord(=CF=87)=20(#17336)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We continue with adding properties of Gauss and Jacobi sums. This PR shows that `g(χ)^n = χ(-1) * #F * J(χ,χ) * J(χ,χ²) * ... * J(χ,χⁿ⁻²)` when `χ` is a multiplicative character on the finite field `F` of order `n` (and with values in a domain). --- Mathlib/NumberTheory/GaussSum.lean | 22 ++++++++++++-- Mathlib/NumberTheory/JacobiSum/Basic.lean | 35 +++++++++++++++++++++++ 2 files changed, 55 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/GaussSum.lean b/Mathlib/NumberTheory/GaussSum.lean index 5140078baa06e..6af4f9545d5ff 100644 --- a/Mathlib/NumberTheory/GaussSum.lean +++ b/Mathlib/NumberTheory/GaussSum.lean @@ -97,8 +97,15 @@ lemma gaussSum_mul {R : Type u} [CommRing R] [Fintype R] {R' : Type v} [CommRing · exact fun a _ ↦ by rw [add_sub_cancel_right, add_comm] rw [sum_congr rfl fun x _ ↦ sum_eq x, sum_comm] --- In the following, we need `R` to be a finite field and `R'` to be a domain. -variable {R : Type u} [Field R] [Fintype R] {R' : Type v} [CommRing R'] [IsDomain R'] +-- In the following, we need `R` to be a finite field. +variable {R : Type u} [Field R] [Fintype R] {R' : Type v} [CommRing R'] + +lemma mul_gaussSum_inv_eq_gaussSum (χ : MulChar R R') (ψ : AddChar R R') : + χ (-1) * gaussSum χ ψ⁻¹ = gaussSum χ ψ := by + rw [ψ.inv_mulShift, ← Units.coe_neg_one] + exact gaussSum_mulShift χ ψ (-1) + +variable [IsDomain R'] -- From now on, `R'` needs to be a domain. -- A helper lemma for `gaussSum_mul_gaussSum_eq_card` below -- Is this useful enough in other contexts to be public? @@ -130,6 +137,17 @@ theorem gaussSum_mul_gaussSum_eq_card {χ : MulChar R R'} (hχ : χ ≠ 1) {ψ : rw [Finset.sum_ite_eq' Finset.univ (1 : R)] simp only [Finset.mem_univ, map_one, one_mul, if_true] +/-- If `χ` is a multiplicative character of order `n` on a finite field `F`, +then `g(χ) * g(χ^(n-1)) = χ(-1)*#F` -/ +lemma gaussSum_mul_gaussSum_pow_orderOf_sub_one {χ : MulChar R R'} {ψ : AddChar R R'} + (hχ : χ ≠ 1) (hψ : ψ.IsPrimitive) : + gaussSum χ ψ * gaussSum (χ ^ (orderOf χ - 1)) ψ = χ (-1) * Fintype.card R := by + have h : χ ^ (orderOf χ - 1) = χ⁻¹ := by + refine (inv_eq_of_mul_eq_one_right ?_).symm + rw [← pow_succ', Nat.sub_one_add_one_eq_of_pos χ.orderOf_pos, pow_orderOf_eq_one] + rw [h, ← mul_gaussSum_inv_eq_gaussSum χ⁻¹, mul_left_comm, gaussSum_mul_gaussSum_eq_card hχ hψ, + MulChar.inv_apply', inv_neg_one] + /-- The Gauss sum of a nontrivial character on a finite field does not vanish. -/ lemma gaussSum_ne_zero_of_nontrivial (h : (Fintype.card R : R') ≠ 0) {χ : MulChar R R'} (hχ : χ ≠ 1) {ψ : AddChar R R'} (hψ : ψ.IsPrimitive) : diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index 1859c852cef47..f9a173578aa98 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -298,3 +298,38 @@ lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ ring end image + +section GaussSum + +variable {F R : Type*} [Fintype F] [Field F] [CommRing R] [IsDomain R] + +lemma gaussSum_pow_eq_prod_jacobiSum_aux (χ : MulChar F R) (ψ : AddChar F R) {n : ℕ} + (hn₁ : 0 < n) (hn₂ : n < orderOf χ) : + gaussSum χ ψ ^ n = gaussSum (χ ^ n) ψ * ∏ j ∈ Ico 1 n, jacobiSum χ (χ ^ j) := by + induction n, hn₁ using Nat.le_induction with + | base => simp only [pow_one, le_refl, Ico_eq_empty_of_le, prod_empty, mul_one] + | succ n hn ih => + specialize ih <| lt_trans (Nat.lt_succ_self n) hn₂ + have gauss_rw : gaussSum (χ ^ n) ψ * gaussSum χ ψ = + jacobiSum χ (χ ^ n) * gaussSum (χ ^ (n + 1)) ψ := by + have hχn : χ * (χ ^ n) ≠ 1 := + pow_succ' χ n ▸ pow_ne_one_of_lt_orderOf n.add_one_ne_zero hn₂ + rw [mul_comm, ← jacobiSum_mul_nontrivial hχn, mul_comm, ← pow_succ'] + apply_fun (· * gaussSum χ ψ) at ih + rw [mul_right_comm, ← pow_succ, gauss_rw] at ih + rw [ih, Finset.prod_Ico_succ_top hn, mul_rotate, mul_assoc] + +/-- If `χ` is a multiplicative character of order `n ≥ 2` on a finite field `F`, +then `g(χ)^n = χ(-1) * #F * J(χ,χ) * J(χ,χ²) * ... * J(χ,χⁿ⁻²)`. -/ +theorem gaussSum_pow_eq_prod_jacobiSum {χ : MulChar F R} {ψ : AddChar F R} (hχ : 2 ≤ orderOf χ) + (hψ : ψ.IsPrimitive) : + gaussSum χ ψ ^ orderOf χ = + χ (-1) * Fintype.card F * ∏ i ∈ Ico 1 (orderOf χ - 1), jacobiSum χ (χ ^ i) := by + have := gaussSum_pow_eq_prod_jacobiSum_aux χ ψ (n := orderOf χ - 1) (by omega) (by omega) + apply_fun (gaussSum χ ψ * ·) at this + rw [← pow_succ', Nat.sub_one_add_one_eq_of_pos (by omega)] at this + have hχ₁ : χ ≠ 1 := + fun h ↦ ((orderOf_one (G := MulChar F R) ▸ h ▸ hχ).trans_lt Nat.one_lt_two).false + rw [this, ← mul_assoc, gaussSum_mul_gaussSum_pow_orderOf_sub_one hχ₁ hψ] + +end GaussSum From f880a608d4dcd56bd79be2644b61e12682a811be Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 2 Oct 2024 05:45:23 +0000 Subject: [PATCH 107/174] docs(Measure/Haar/OfBasis): expand a docstring (#17270) --- Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 5c2649473b452..18ebba51fd6f0 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -290,7 +290,18 @@ end Fintype /-- A finite dimensional inner product space has a canonical measure, the Lebesgue measure giving volume `1` to the parallelepiped spanned by any orthonormal basis. We define the measure using some arbitrary choice of orthonormal basis. The fact that it works with any orthonormal basis -is proved in `orthonormalBasis.volume_parallelepiped`. -/ +is proved in `orthonormalBasis.volume_parallelepiped`. + +This instance creates: + +- a potential non-defeq diamond with the natural instance for `MeasureSpace (ULift E)`, + which does not exist in Mathlib at the moment; + +- a diamond with the existing instance `MeasureTheory.Measure.instMeasureSpacePUnit`. + +However, we've decided not to refactor until one of these diamonds starts creating issues, see +https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Hausdorff.20measure.20normalisation +-/ instance (priority := 100) measureSpaceOfInnerProductSpace [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] : MeasureSpace E where volume := (stdOrthonormalBasis ℝ E).toBasis.addHaar From d902abe2c36d37eb76fddd9e43290bcef3bcb6e8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 2 Oct 2024 05:45:24 +0000 Subject: [PATCH 108/174] chore: make a global instance local in ModelTheory.Order (#17271) --- Mathlib/ModelTheory/Order.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/ModelTheory/Order.lean b/Mathlib/ModelTheory/Order.lean index ac41879019da5..af11c905aa569 100644 --- a/Mathlib/ModelTheory/Order.lean +++ b/Mathlib/ModelTheory/Order.lean @@ -358,10 +358,14 @@ instance : @OrderedStructure L M _ (L.leOfStructure M) _ := by intros rfl -instance [h : DecidableRel (fun (a b : M) => Structure.RelMap (leSymb : L.Relations 2) ![a,b])] : - DecidableRel (@LE.le M (L.leOfStructure M)) := by - letI := L.leOfStructure M - exact h +/-- The order structure on an ordered language is decidable. -/ +-- This should not be a global instance, +-- because it will match with any `LE` typeclass search +@[local instance] +def decidableLEOfStructure + [h : DecidableRel (fun (a b : M) => Structure.RelMap (leSymb : L.Relations 2) ![a,b])] : + letI := L.leOfStructure M + DecidableRel ((· : M) ≤ ·) := h /-- Any model of a theory of preorders is a preorder. -/ def preorderOfModels [h : M ⊨ L.preorderTheory] : Preorder M where From 3fdb944d51949cc06cd938cdf78df3e87ef4c5c3 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 2 Oct 2024 06:17:33 +0000 Subject: [PATCH 109/174] chore: prepare for multi-goal linter (#17316) Adaptations for #12339. --- Mathlib/Computability/TMToPartrec.lean | 6 +++--- Mathlib/Data/Seq/Computation.lean | 4 ++-- Mathlib/Order/Defs.lean | 8 ++++---- 3 files changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 68249b2638fe3..79a322e54c754 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1604,7 +1604,7 @@ def trStmts₁ : Λ' → Finset Λ' theorem trStmts₁_trans {q q'} : q' ∈ trStmts₁ q → trStmts₁ q' ⊆ trStmts₁ q := by induction q with | move _ _ _ q q_ih => _ | clear _ _ q q_ih => _ | copy q q_ih => _ | push _ _ q q_ih => _ - | read q q_ih => _ | succ q q_ih => _ | pred q₁ q₂ q₁_ih q₂_ih => _ | ret => _ + | read q q_ih => _ | succ q q_ih => _ | pred q₁ q₂ q₁_ih q₂_ih => _ | ret => _ <;> all_goals simp (config := { contextual := true }) only [trStmts₁, Finset.mem_insert, Finset.mem_union, or_imp, Finset.mem_singleton, Finset.Subset.refl, imp_true_iff, true_and] @@ -1801,8 +1801,8 @@ theorem trStmts₁_supports {S q} (H₁ : (q : Λ').Supports S) (HS₁ : trStmts have W := fun {q} => trStmts₁_self q induction q with | move _ _ _ q q_ih => _ | clear _ _ q q_ih => _ | copy q q_ih => _ | push _ _ q q_ih => _ - | read q q_ih => _ | succ q q_ih => _ | pred q₁ q₂ q₁_ih q₂_ih => _ | ret => _ - all_goals simp [trStmts₁, -Finset.singleton_subset_iff] at HS₁ ⊢ + | read q q_ih => _ | succ q q_ih => _ | pred q₁ q₂ q₁_ih q₂_ih => _ | ret => _ <;> + simp [trStmts₁, -Finset.singleton_subset_iff] at HS₁ ⊢ any_goals cases' Finset.insert_subset_iff.1 HS₁ with h₁ h₂ first | have h₃ := h₂ W | try simp [Finset.subset_iff] at h₂ diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index 4622e5d162c96..9ff96e3dd47ef 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -503,8 +503,8 @@ theorem length_thinkN (s : Computation α) [_h : Terminates s] (n) : theorem eq_thinkN {s : Computation α} {a n} (h : Results s a n) : s = thinkN (pure a) n := by revert s - induction n with | zero => _ | succ n IH => _ - all_goals intro s; apply recOn s (fun a' => _) fun s => _ <;> intro a h + induction n with | zero => _ | succ n IH => _ <;> + (intro s; apply recOn s (fun a' => _) fun s => _) <;> intro a h · rw [← eq_of_pure_mem h.mem] rfl · cases' of_results_think h with n h diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index d50b65de7d6d3..1bdaefb9a2491 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -501,9 +501,9 @@ lemma min_comm (a b : α) : min a b = min b a := lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by apply eq_min - · apply le_trans; apply min_le_left; apply min_le_left + · apply le_trans (min_le_left ..); apply min_le_left · apply le_min - · apply le_trans; apply min_le_left; apply min_le_right + · apply le_trans (min_le_left ..); apply min_le_right · apply min_le_right · intro d h₁ h₂; apply le_min · apply le_min h₁; apply le_trans h₂; apply min_le_left @@ -528,9 +528,9 @@ lemma max_comm (a b : α) : max a b = max b a := lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by apply eq_max - · apply le_trans; apply le_max_left a b; apply le_max_left + · apply le_trans (le_max_left a b); apply le_max_left · apply max_le - · apply le_trans; apply le_max_right a b; apply le_max_left + · apply le_trans (le_max_right a b); apply le_max_left · apply le_max_right · intro d h₁ h₂; apply max_le · apply max_le h₁; apply le_trans (le_max_left _ _) h₂ From b5fa3ef08e3ede6b7a4bb1795bd0fd43fd387468 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 2 Oct 2024 07:21:32 +0000 Subject: [PATCH 110/174] CI: add high import percentage changes (#17343) This PR adds automation for detecting PRs with a significant increase in transitive imports. Significant means a single file with a percentage increase of at least 2%. If there are such files, then the automation adds * a table with the import-increased files at the beginning of the `PR summary`; * the `large-import` label. If there are no such files, then the bot removes the `large-import` label, if it was present. The PR summary of this PR shows that there is no change when there are no significant changes in transitive imports. #17344 shows what happens when such a change is present. --- .github/workflows/PR_summary.yml | 19 +++++++++++++++++-- scripts/import-graph-report.py | 17 +++++++++++++++-- 2 files changed, 32 insertions(+), 4 deletions(-) diff --git a/.github/workflows/PR_summary.yml b/.github/workflows/PR_summary.yml index 25529d869f64b..bbde4ace8c7e5 100644 --- a/.github/workflows/PR_summary.yml +++ b/.github/workflows/PR_summary.yml @@ -55,12 +55,27 @@ jobs: PR="${{ github.event.pull_request.number }}" title="### PR summary" + graphAndHighPercentReports=$(python ./scripts/import-graph-report.py base.json head.json changed_files.txt) + ## Import count comment importCount=$( - python ./scripts/import-graph-report.py base.json head.json changed_files.txt + printf '%s\n' "${graphAndHighPercentReports}" | sed '/^Import changes exceeding/Q' ./scripts/import_trans_difference.sh ) + ## High percentage imports + high_percentages=$( + printf '%s\n' "${graphAndHighPercentReports}" | sed -n '/^Import changes exceeding/,$p' + ) + # if there are files with large increase in transitive imports, then we add the `large-import` label + if [ -n "${high_percentages}" ] + then + high_percentages=$'\n\n'"${high_percentages}" + gh pr edit "${PR}" --add-label large-import + else # otherwise, we remove the label + gh pr edit "${PR}" --remove-label large-import + fi + if [ "$(printf '%s' "${importCount}" | wc -l)" -gt 12 ] then importCount="$(printf '
\n\n%s\n\n\n\n%s\n\n
\n' "#### Import changes for modified files" "${importCount}")" @@ -80,6 +95,6 @@ jobs: currentHash="$(git rev-parse HEAD)" hashURL="https://github.com/${{ github.repository }}/pull/${{ github.event.pull_request.number }}/commits/${currentHash}" - message="$(printf '%s [%s](%s)\n\n%s\n\n---\n\n%s\n' "${title}" "$(git rev-parse --short HEAD)" "${hashURL}" "${importCount}" "${declDiff}")" + message="$(printf '%s [%s](%s)%s\n\n%s\n\n---\n\n%s\n' "${title}" "$(git rev-parse --short HEAD)" "${hashURL}" "${high_percentages}" "${importCount}" "${declDiff}")" ./scripts/update_PR_comment.sh "${message}" "${title}" "${PR}" diff --git a/scripts/import-graph-report.py b/scripts/import-graph-report.py index cf072b916c5e8..86775bf4c2fc1 100755 --- a/scripts/import-graph-report.py +++ b/scripts/import-graph-report.py @@ -10,6 +10,8 @@ import json import sys +high_import_threshold = 2 + def compare_counts(base_file, head_file, changed_files_txt): # Load the counts with open(head_file, 'r') as f: @@ -29,6 +31,7 @@ def compare_counts(base_file, head_file, changed_files_txt): # Compare the counts changes = [] + high_pct = [] for file in changed_files: base_count = base_counts.get(file, 0) head_count = head_counts.get(file, 0) @@ -36,6 +39,8 @@ def compare_counts(base_file, head_file, changed_files_txt): continue diff = head_count - base_count percent = (diff / base_count) * 100 + if high_import_threshold < percent: + high_pct.append(f'| +{percent:.2f}% | `{file}` |') if diff < 0: # Dependencies went down changes.append((file, base_count, head_count, diff, percent)) elif diff > new_files: # Dependencies went up by more than the number of new files @@ -59,11 +64,19 @@ def compare_counts(base_file, head_file, changed_files_txt): message += '\n'.join(messages) else: message += 'No significant changes to the import graph' - return message + + high_pct_report = '' + if high_pct: + high_pct_report += f'Import changes exceeding {high_import_threshold}%\n\n' + high_pct_report += '| % | File |\n' + high_pct_report += '| - | - |\n' + high_pct_report += '\n'.join(high_pct) + return (message, high_pct_report) if __name__ == '__main__': base_file = sys.argv[1] head_file = sys.argv[2] changed_files_txt = sys.argv[3] - message = compare_counts(base_file, head_file, changed_files_txt) + (message, high_pct) = compare_counts(base_file, head_file, changed_files_txt) print(message) + print(high_pct) From f27f2322c8ce6c37b3ffed293f2c212e2852acbf Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Wed, 2 Oct 2024 08:44:09 +0000 Subject: [PATCH 111/174] feat(CategoryTheory/Adjunction): composing a right adjoint with yoneda (#17333) --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index f0a274f9be2a4..30d3ab3bb1913 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton, Johan Commelin, Bhavik Mehta -/ import Mathlib.CategoryTheory.Equivalence +import Mathlib.CategoryTheory.Yoneda /-! # Adjunctions between functors @@ -419,6 +420,22 @@ def ofNatIsoRight {F : C ⥤ D} {G H : D ⥤ C} (adj : F ⊣ G) (iso : G ≅ H) Adjunction.mkOfHomEquiv { homEquiv := fun X Y => (adj.homEquiv X Y).trans (equivHomsetRightOfNatIso iso) } +/-- The isomorpism which an adjunction `F ⊣ G` induces on `G ⋙ yoneda`. This states that +`Adjunction.homEquiv` is natural in both arguments. -/ +@[simps!] +def compYonedaIso {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₁} D] + {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : + G ⋙ yoneda ≅ yoneda ⋙ (whiskeringLeft _ _ _).obj F.op := + NatIso.ofComponents fun X => NatIso.ofComponents fun Y => (adj.homEquiv Y.unop X).toIso.symm + +/-- The isomorpism which an adjunction `F ⊣ G` induces on `F.op ⋙ coyoneda`. This states that +`Adjunction.homEquiv` is natural in both arguments. -/ +@[simps!] +def compCoyonedaIso {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₁} D] + {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : + F.op ⋙ coyoneda ≅ coyoneda ⋙ (whiskeringLeft _ _ _).obj G := + NatIso.ofComponents fun X => NatIso.ofComponents fun Y => (adj.homEquiv X.unop Y).toIso + section variable {E : Type u₃} [ℰ : Category.{v₃} E] {H : D ⥤ E} {I : E ⥤ D} From e4a8f5b0444a8dcd89f1ac4b7602f894bb177516 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 09:10:58 +0000 Subject: [PATCH 112/174] feat: balancing a function (#17205) Define the balancing of a function, namely the function minus its average. Doing this operation has the effect of nullifying the `0`-th component of the Fourier transform (but this PR doesn't prove this fact). From LeanAPAP Co-authored-by: Bhavik Mehta --- Mathlib.lean | 1 + Mathlib/Algebra/BigOperators/Balance.lean | 56 +++++++++++++++++++++++ 2 files changed, 57 insertions(+) create mode 100644 Mathlib/Algebra/BigOperators/Balance.lean diff --git a/Mathlib.lean b/Mathlib.lean index 96d5225a538f8..24ae2e8019f4b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -37,6 +37,7 @@ import Mathlib.Algebra.AlgebraicCard import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.Associated.OrderedCommMonoid import Mathlib.Algebra.BigOperators.Associated +import Mathlib.Algebra.BigOperators.Balance import Mathlib.Algebra.BigOperators.Expect import Mathlib.Algebra.BigOperators.Fin import Mathlib.Algebra.BigOperators.Finprod diff --git a/Mathlib/Algebra/BigOperators/Balance.lean b/Mathlib/Algebra/BigOperators/Balance.lean new file mode 100644 index 0000000000000..1b2b2ca767a64 --- /dev/null +++ b/Mathlib/Algebra/BigOperators/Balance.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2023 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.BigOperators.Expect + +/-! +# Balancing a function + +This file defines the balancing of a function `f`, defined as `f` minus its average. + +This is the unique function `g` such that `f a - f b = g a - g b` for all `a` and `b`, and +`∑ a, g a = 0`. This is particularly useful in Fourier analysis as `f` and `g` then have the same +Fourier transform, except in the `0`-th frequency where the Fourier transform of `g` vanishes. +-/ + +open Finset Function +open scoped BigOperators + +variable {ι H F G : Type*} + +namespace Fintype + +section AddCommGroup +variable [Fintype ι] [AddCommGroup G] [Module ℚ≥0 G] [AddCommGroup H] [Module ℚ≥0 H] + +/-- The balancing of a function, namely the function minus its average. -/ +def balance (f : ι → G) : ι → G := f - Function.const _ (𝔼 y, f y) + +lemma balance_apply (f : ι → G) (x : ι) : balance f x = f x - 𝔼 y, f y := rfl + +@[simp] lemma balance_zero : balance (0 : ι → G) = 0 := by simp [balance] + +@[simp] lemma balance_add (f g : ι → G) : balance (f + g) = balance f + balance g := by + simp only [balance, expect_add_distrib, ← const_add, add_sub_add_comm, Pi.add_apply] + +@[simp] lemma balance_sub (f g : ι → G) : balance (f - g) = balance f - balance g := by + simp only [balance, expect_sub_distrib, const_sub, sub_sub_sub_comm, Pi.sub_apply] + +@[simp] lemma balance_neg (f : ι → G) : balance (-f) = -balance f := by + simp only [balance, expect_neg_distrib, const_neg, neg_sub', Pi.neg_apply] + +@[simp] lemma sum_balance (f : ι → G) : ∑ x, balance f x = 0 := by + cases isEmpty_or_nonempty ι <;> simp [balance_apply] + +@[simp] lemma expect_balance (f : ι → G) : 𝔼 x, balance f x = 0 := by simp [expect] + +@[simp] lemma balance_idem (f : ι → G) : balance (balance f) = balance f := by + cases isEmpty_or_nonempty ι <;> ext x <;> simp [balance, expect_sub_distrib, univ_nonempty] + +@[simp] lemma map_balance [FunLike F G H] [LinearMapClass F ℚ≥0 G H] (g : F) (f : ι → G) (a : ι) : + g (balance f a) = balance (g ∘ f) a := by simp [balance, map_expect] + +end AddCommGroup +end Fintype From 7fc7838ae4893e5ae499695841fe5c57ee4c0a08 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Wed, 2 Oct 2024 09:20:04 +0000 Subject: [PATCH 113/174] feat(CategoryTheory/Biproducts): when a category has biproducts, lim and colim functors are isomorphic (#17104) --- .../Limits/Shapes/Biproducts.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index dbb941df4e33f..c5ab7690aefbc 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -403,6 +403,14 @@ instance (priority := 100) hasFiniteCoproducts_of_hasFiniteBiproducts [HasFinite HasFiniteCoproducts C where out _ := ⟨fun _ => hasColimitOfIso Discrete.natIsoFunctor⟩ +instance (priority := 100) hasProductsOfShape_of_hasBiproductsOfShape [HasBiproductsOfShape J C] : + HasProductsOfShape J C where + has_limit _ := hasLimitOfIso Discrete.natIsoFunctor.symm + +instance (priority := 100) hasCoproductsOfShape_of_hasBiproductsOfShape [HasBiproductsOfShape J C] : + HasCoproductsOfShape J C where + has_colimit _ := hasColimitOfIso Discrete.natIsoFunctor + variable {C} /-- The isomorphism between the specified limit and the specified colimit for @@ -547,6 +555,17 @@ theorem biproduct.isoCoproduct_hom {f : J → C} [HasBiproduct f] : (biproduct.isoCoproduct f).hom = biproduct.desc (Sigma.ι f) := biproduct.hom_ext' _ _ fun j => by simp [← Iso.eq_comp_inv] +/-- If a category has biproducts of a shape `J`, its `colim` and `lim` functor on diagrams over `J` +are isomorphic. -/ +@[simps!] +def HasBiproductsOfShape.colimIsoLim [HasBiproductsOfShape J C] : + colim (J := Discrete J) (C := C) ≅ lim := + NatIso.ofComponents (fun F => (Sigma.isoColimit F).symm ≪≫ + (biproduct.isoCoproduct _).symm ≪≫ biproduct.isoProduct _ ≪≫ Pi.isoLimit F) + fun η => colimit.hom_ext fun ⟨i⟩ => limit.hom_ext fun ⟨j⟩ => by + by_cases h : i = j <;> + simp_all [h, Sigma.isoColimit, Pi.isoLimit, biproduct.ι_π, biproduct.ι_π_assoc] + theorem biproduct.map_eq_map' {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ b, f b ⟶ g b) : biproduct.map p = biproduct.map' p := by ext From b91c5b63fd9ad104b16ea34bd259162dd1850b9e Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 2 Oct 2024 09:44:47 +0000 Subject: [PATCH 114/174] feat(CategoryTheory/Galois): fiber functor is essentially surjective (#16552) We show that every finite, discrete `Aut F`-set is in the essential image of the fiber functor. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Galois/EssSurj.lean | 263 +++++++++++++++++++++ Mathlib/CategoryTheory/Galois/Full.lean | 2 +- 3 files changed, 265 insertions(+), 1 deletion(-) create mode 100644 Mathlib/CategoryTheory/Galois/EssSurj.lean diff --git a/Mathlib.lean b/Mathlib.lean index 24ae2e8019f4b..2851c53609ce6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1542,6 +1542,7 @@ import Mathlib.CategoryTheory.Functor.Trifunctor import Mathlib.CategoryTheory.Galois.Action import Mathlib.CategoryTheory.Galois.Basic import Mathlib.CategoryTheory.Galois.Decomposition +import Mathlib.CategoryTheory.Galois.EssSurj import Mathlib.CategoryTheory.Galois.Examples import Mathlib.CategoryTheory.Galois.Full import Mathlib.CategoryTheory.Galois.GaloisObjects diff --git a/Mathlib/CategoryTheory/Galois/EssSurj.lean b/Mathlib/CategoryTheory/Galois/EssSurj.lean new file mode 100644 index 0000000000000..06f9db4d4af8d --- /dev/null +++ b/Mathlib/CategoryTheory/Galois/EssSurj.lean @@ -0,0 +1,263 @@ +/- +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.Galois.Full +import Mathlib.CategoryTheory.Galois.Topology +import Mathlib.Topology.Algebra.OpenSubgroup + +/-! + +# Essential surjectivity of fiber functors + +Let `F : C ⥤ FintypeCat` be a fiber functor of a Galois category `C` and denote by +`H` the induced functor `C ⥤ Action FintypeCat (Aut F)`. + +In this file we show that the essential image of `H` consists of the finite `Aut F`-sets where +the `Aut F` action is continuous. + +## Main results + +- `exists_lift_of_quotient_openSubgroup`: If `U` is an open subgroup of `Aut F`, then + there exists an object `X` such that `F.obj X` is isomorphic to `Aut F ⧸ U` as + `Aut F`-sets. +- `exists_lift_of_continuous`: If `X` is a finite, discrete `Aut F`-set, then + there exists an object `A` such that `F.obj A` is isomorphic to `X` as + `Aut F`-sets. + +## Strategy + +We first show that every finite, discrete `Aut F`-set `Y` has a decomposition into connected +components and each connected component is of the form `Aut F ⧸ U` for an open subgroup `U`. +Since `H` preserves finite coproducts, it hence suffices to treat the case `Y = Aut F ⧸ U`. +For the case `Y = Aut F ⧸ U` we closely follow the second part of Stacks Project Tag 0BN4. + +-/ + +noncomputable section + +universe u₁ u₂ + +namespace CategoryTheory + +namespace PreGaloisCategory + +variable {C : Type u₁} [Category.{u₂} C] {F : C ⥤ FintypeCat.{u₁}} + +open Limits Functor + +variable [GaloisCategory C] [FiberFunctor F] + +variable {G : Type*} [Group G] [TopologicalSpace G] [TopologicalGroup G] [CompactSpace G] + +private local instance fintypeQuotient (H : OpenSubgroup (G)) : + Fintype (G ⧸ (H : Subgroup (G))) := + have : Finite (G ⧸ H.toSubgroup) := H.toSubgroup.quotient_finite_of_isOpen H.isOpen' + Fintype.ofFinite _ + +private local instance fintypeQuotientStabilizer {X : Type*} [MulAction G X] + [TopologicalSpace X] [ContinuousSMul G X] [DiscreteTopology X] (x : X) : + Fintype (G ⧸ (MulAction.stabilizer (G) x)) := + fintypeQuotient ⟨MulAction.stabilizer (G) x, stabilizer_isOpen (G) x⟩ + +/-- If `X` is a finite discrete `G`-set, it can be written as the finite disjoint union +of quotients of the form `G ⧸ Uᵢ` for open subgroups `(Uᵢ)`. Note that this +is simply the decomposition into orbits. -/ +lemma has_decomp_quotients (X : Action FintypeCat (MonCat.of G)) + [TopologicalSpace X.V] [DiscreteTopology X.V] [ContinuousSMul G X.V] : + ∃ (ι : Type) (_ : Finite ι) (f : ι → OpenSubgroup (G)), + Nonempty ((∐ fun i ↦ G ⧸ₐ (f i).toSubgroup) ≅ X) := by + obtain ⟨ι, hf, f, u, hc⟩ := has_decomp_connected_components' X + letI (i : ι) : TopologicalSpace (f i).V := ⊥ + haveI (i : ι) : DiscreteTopology (f i).V := ⟨rfl⟩ + have (i : ι) : ContinuousSMul G (f i).V := ContinuousSMul.mk <| by + let r : f i ⟶ X := Sigma.ι f i ≫ u.hom + let r'' (p : G × (f i).V) : G × X.V := (p.1, r.hom p.2) + let q (p : G × X.V) : X.V := X.ρ p.1 p.2 + let q' (p : G × (f i).V) : (f i).V := (f i).ρ p.1 p.2 + have heq : q ∘ r'' = r.hom ∘ q' := by + ext (p : G × (f i).V) + exact (congr_fun (r.comm p.1) p.2).symm + have hrinj : Function.Injective r.hom := + (ConcreteCategory.mono_iff_injective_of_preservesPullback r).mp <| mono_comp _ _ + let t₁ : TopologicalSpace (G × (f i).V) := inferInstance + show @Continuous _ _ _ ⊥ q' + have : TopologicalSpace.induced r.hom inferInstance = ⊥ := by + rw [← le_bot_iff] + exact fun s _ ↦ ⟨r.hom '' s, ⟨isOpen_discrete (r.hom '' s), Set.preimage_image_eq s hrinj⟩⟩ + rw [← this, continuous_induced_rng, ← heq] + exact Continuous.comp continuous_smul (by fun_prop) + have (i : ι) : ∃ (U : OpenSubgroup (G)), (Nonempty ((f i) ≅ G ⧸ₐ U.toSubgroup)) := by + obtain ⟨(x : (f i).V)⟩ := nonempty_fiber_of_isConnected (forget₂ _ _) (f i) + let U : OpenSubgroup (G) := ⟨MulAction.stabilizer (G) x, stabilizer_isOpen (G) x⟩ + letI : Fintype (G ⧸ MulAction.stabilizer (G) x) := fintypeQuotient U + exact ⟨U, ⟨FintypeCat.isoQuotientStabilizerOfIsConnected (f i) x⟩⟩ + choose g ui using this + exact ⟨ι, hf, g, ⟨(Sigma.mapIso (fun i ↦ (ui i).some)).symm ≪≫ u⟩⟩ + +/-- If `X` is connected and `x` is in the fiber of `X`, `F.obj X` is isomorphic +to the quotient of `Aut F` by the stabilizer of `x` as `Aut F`-sets. -/ +def fiberIsoQuotientStabilizer (X : C) [IsConnected X] (x : F.obj X) : + (functorToAction F).obj X ≅ Aut F ⧸ₐ MulAction.stabilizer (Aut F) x := + haveI : IsConnected ((functorToAction F).obj X) := PreservesIsConnected.preserves + letI : Fintype (Aut F ⧸ MulAction.stabilizer (Aut F) x) := fintypeQuotientStabilizer x + FintypeCat.isoQuotientStabilizerOfIsConnected ((functorToAction F).obj X) x + +section + +open Action.FintypeCat + +variable (V : OpenSubgroup (Aut F)) {U : OpenSubgroup (Aut F)} + (h : Subgroup.Normal U.toSubgroup) {A : C} (u : (functorToAction F).obj A ≅ Aut F ⧸ₐ U.toSubgroup) + +/- + +### Strategy outline + +Let `A` be an object of `C` with fiber `Aut F`-isomorphic to `Aut F ⧸ U` for an open normal +subgroup `U`. Then for any open subgroup `V` of `Aut F`, `V ⧸ (U ⊓ V)` acts on `A`. This +induces the diagram `quotientDiag`. Now assume `U ≤ V`. Then we can also postcompose +the diagram `quotientDiag` with `F`. The goal of this section is to compute that the colimit +of this composed diagram is `Aut F ⧸ V`. Finally, we obtain `F.obj (A ⧸ V) ≅ Aut F ⧸ V` as +`Aut F`-sets. +-/ + +private def quotientToEndObjectHom : + V.toSubgroup ⧸ Subgroup.subgroupOf U.toSubgroup V.toSubgroup →* End A := + let ff : (functorToAction F).FullyFaithful := FullyFaithful.ofFullyFaithful (functorToAction F) + let e : End A ≃* End (Aut F ⧸ₐ U.toSubgroup) := (ff.mulEquivEnd A).trans (Iso.conj u) + e.symm.toMonoidHom.comp (quotientToEndHom V.toSubgroup U.toSubgroup) + +private lemma functorToAction_map_quotientToEndObjectHom + (m : SingleObj.star (V ⧸ Subgroup.subgroupOf U.toSubgroup V.toSubgroup) ⟶ + SingleObj.star (V ⧸ Subgroup.subgroupOf U.toSubgroup V.toSubgroup)) : + (functorToAction F).map (quotientToEndObjectHom V h u m) = + u.hom ≫ quotientToEndHom V.toSubgroup U.toSubgroup m ≫ u.inv := by + simp [← cancel_epi u.inv, ← cancel_mono u.hom, ← Iso.conj_apply, quotientToEndObjectHom] + +@[simps!] +private def quotientDiag : SingleObj (V.toSubgroup ⧸ Subgroup.subgroupOf U V) ⥤ C := + SingleObj.functor (quotientToEndObjectHom V h u) + +variable {V} (hUinV : U ≤ V) + +@[simps] +private def coconeQuotientDiag : + Cocone (quotientDiag V h u ⋙ functorToAction F) where + pt := Aut F ⧸ₐ V.toSubgroup + ι := SingleObj.natTrans (u.hom ≫ quotientToQuotientOfLE V.toSubgroup U.toSubgroup hUinV) <| by + intro (m : V ⧸ Subgroup.subgroupOf U V) + simp only [const_obj_obj, Functor.comp_map, const_obj_map, Category.comp_id] + rw [← cancel_epi (u.inv), Iso.inv_hom_id_assoc] + apply Action.hom_ext + ext (x : Aut F ⧸ U.toSubgroup) + induction' m, x using Quotient.inductionOn₂ with σ μ + suffices h : ⟦μ * σ⁻¹⟧ = ⟦μ⟧ by + simp only [quotientToQuotientOfLE_hom_mk, quotientDiag_map, + functorToAction_map_quotientToEndObjectHom V _ u] + simpa + apply Quotient.sound + apply (QuotientGroup.leftRel_apply).mpr + simp + +@[simps] +private def coconeQuotientDiagDesc + (s : Cocone (quotientDiag V h u ⋙ functorToAction F)) : + (coconeQuotientDiag h u hUinV).pt ⟶ s.pt where + hom := Quotient.lift (fun σ ↦ (u.inv ≫ s.ι.app (SingleObj.star _)).hom ⟦σ⟧) <| fun σ τ hst ↦ by + let J' := quotientDiag V h u ⋙ functorToAction F + let m : End (SingleObj.star (V.toSubgroup ⧸ Subgroup.subgroupOf U V)) := + ⟦⟨σ⁻¹ * τ, (QuotientGroup.leftRel_apply).mp hst⟩⟧ + have h1 : J'.map m ≫ s.ι.app (SingleObj.star _) = s.ι.app (SingleObj.star _) := s.ι.naturality m + conv_rhs => rw [← h1] + have h2 : (J'.map m).hom (u.inv.hom ⟦τ⟧) = u.inv.hom ⟦σ⟧ := by + simp only [comp_obj, quotientDiag_obj, Functor.comp_map, quotientDiag_map, J', + functorToAction_map_quotientToEndObjectHom V h u m] + show (u.inv ≫ u.hom ≫ _ ≫ u.inv).hom ⟦τ⟧ = u.inv.hom ⟦σ⟧ + simp [m] + simp only [← h2, const_obj_obj, Action.comp_hom, FintypeCat.comp_apply] + comm g := by + ext (x : Aut F ⧸ V.toSubgroup) + induction' x using Quotient.inductionOn with σ + simp only [const_obj_obj] + show (((Aut F ⧸ₐ U.toSubgroup).ρ g ≫ u.inv.hom) ≫ (s.ι.app (SingleObj.star _)).hom) ⟦σ⟧ = + ((s.ι.app (SingleObj.star _)).hom ≫ s.pt.ρ g) (u.inv.hom ⟦σ⟧) + have : ((functorToAction F).obj A).ρ g ≫ (s.ι.app (SingleObj.star _)).hom = + (s.ι.app (SingleObj.star _)).hom ≫ s.pt.ρ g := + (s.ι.app (SingleObj.star _)).comm g + rw [← this, u.inv.comm g] + rfl + +/-- The constructed cocone `coconeQuotientDiag` on the diagram `quotientDiag` is colimiting. -/ +private def coconeQuotientDiagIsColimit : + IsColimit (coconeQuotientDiag h u hUinV) where + desc := coconeQuotientDiagDesc h u hUinV + fac s j := by + apply (cancel_epi u.inv).mp + apply Action.hom_ext + ext (x : Aut F ⧸ U.toSubgroup) + induction' x using Quotient.inductionOn with σ + simp + rfl + uniq s f hf := by + apply Action.hom_ext + ext (x : Aut F ⧸ V.toSubgroup) + induction' x using Quotient.inductionOn with σ + simp [← hf (SingleObj.star _)] + +end + +/-- For every open subgroup `V` of `Aut F`, there exists an `X : C` such that +`F.obj X ≅ Aut F ⧸ V` as `Aut F`-sets. -/ +lemma exists_lift_of_quotient_openSubgroup (V : OpenSubgroup (Aut F)) : + ∃ (X : C), Nonempty ((functorToAction F).obj X ≅ Aut F ⧸ₐ V.toSubgroup) := by + obtain ⟨I, hf, hc, hi⟩ := exists_set_ker_evaluation_subset_of_isOpen F (one_mem V) V.isOpen' + haveI (X : I) : IsConnected X.val := hc X X.property + haveI (X : I) : Nonempty (F.obj X.val) := nonempty_fiber_of_isConnected F X + have hn : Nonempty (F.obj <| (∏ᶜ fun X : I => X)) := nonempty_fiber_pi_of_nonempty_of_finite F _ + obtain ⟨A, f, hgal⟩ := exists_hom_from_galois_of_fiber_nonempty F (∏ᶜ fun X : I => X) hn + obtain ⟨a⟩ := nonempty_fiber_of_isConnected F A + let U : OpenSubgroup (Aut F) := ⟨MulAction.stabilizer (Aut F) a, stabilizer_isOpen (Aut F) a⟩ + let u := fiberIsoQuotientStabilizer A a + have hUnormal : U.toSubgroup.Normal := stabilizer_normal_of_isGalois F A a + have h1 (σ : Aut F) (σinU : σ ∈ U) : σ.hom.app A = 𝟙 (F.obj A) := by + have hi : (Aut F ⧸ₐ MulAction.stabilizer (Aut F) a).ρ σ = 𝟙 _ := by + refine FintypeCat.hom_ext _ _ (fun x ↦ ?_) + induction' x using Quotient.inductionOn with τ + show ⟦σ * τ⟧ = ⟦τ⟧ + apply Quotient.sound + apply (QuotientGroup.leftRel_apply).mpr + simp only [mul_inv_rev] + exact Subgroup.Normal.conj_mem hUnormal _ (Subgroup.inv_mem U.toSubgroup σinU) _ + simp [← cancel_mono u.hom.hom, show σ.hom.app A ≫ u.hom.hom = _ from u.hom.comm σ, hi] + have h2 (σ : Aut F) (σinU : σ ∈ U) : ∀ X : I, σ.hom.app X = 𝟙 (F.obj X) := by + intro ⟨X, hX⟩ + ext (x : F.obj X) + let p : A ⟶ X := f ≫ Pi.π (fun Z : I => (Z : C)) ⟨X, hX⟩ + have : IsConnected X := hc X hX + obtain ⟨a, rfl⟩ := surjective_of_nonempty_fiber_of_isConnected F p x + simp only [FintypeCat.id_apply, FunctorToFintypeCat.naturality, h1 σ σinU] + have hUinV : (U : Set (Aut F)) ≤ V := fun u uinU ↦ hi u (h2 u uinU) + have := V.quotient_finite_of_isOpen' (U.subgroupOf V) V.isOpen (V.subgroupOf_isOpen U U.isOpen) + exact ⟨colimit (quotientDiag V hUnormal u), + ⟨preservesColimitIso (functorToAction F) (quotientDiag V hUnormal u) ≪≫ + colimit.isoColimitCocone ⟨coconeQuotientDiag hUnormal u hUinV, + coconeQuotientDiagIsColimit hUnormal u hUinV⟩⟩⟩ + +/-- +If `X` is a finite, discrete `Aut F`-set with continuous `Aut F`-action, then +there exists `A : C` such that `F.obj A ≅ X` as `Aut F`-sets. +-/ +@[stacks 0BN4 "Essential surjectivity part"] +theorem exists_lift_of_continuous (X : Action FintypeCat (MonCat.of (Aut F))) + [TopologicalSpace X.V] [DiscreteTopology X.V] [ContinuousSMul (Aut F) X.V] : + ∃ A, Nonempty ((functorToAction F).obj A ≅ X) := by + obtain ⟨ι, hfin, f, ⟨u⟩⟩ := has_decomp_quotients X + choose g gu using (fun i ↦ exists_lift_of_quotient_openSubgroup (f i)) + exact ⟨∐ g, ⟨PreservesCoproduct.iso (functorToAction F) g ≪≫ + Sigma.mapIso (fun i ↦ (gu i).some) ≪≫ u⟩⟩ + +end PreGaloisCategory + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Galois/Full.lean b/Mathlib/CategoryTheory/Galois/Full.lean index db7f5c8139acc..4c09add2e62f9 100644 --- a/Mathlib/CategoryTheory/Galois/Full.lean +++ b/Mathlib/CategoryTheory/Galois/Full.lean @@ -42,7 +42,7 @@ namespace PreGaloisCategory open Limits Functor -variable {C : Type u} [Category.{u} C] (F : C ⥤ FintypeCat.{u}) [GaloisCategory C] [FiberFunctor F] +variable {C : Type u} [Category.{v} C] (F : C ⥤ FintypeCat.{u}) [GaloisCategory C] [FiberFunctor F] /-- Let `X` be an object of a Galois category with fiber functor `F` and `Y` a sub-`Aut F`-set From a33739dc9104bfa7b4f636b3c80803031b184ee9 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 2 Oct 2024 10:44:40 +0000 Subject: [PATCH 115/174] feat: add relation between subset and + for multisets (#17345) Naming convention and argument explicitness follows the `Set.subset_union_**` lemmas. --- Mathlib/Data/Multiset/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 9cd7fc8586622..c124885555f3b 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -601,6 +601,10 @@ theorem le_add_right (s t : Multiset α) : s ≤ s + t := by simpa using add_le_ theorem le_add_left (s t : Multiset α) : s ≤ t + s := by simpa using add_le_add_right (zero_le t) s +lemma subset_add_left {s t : Multiset α} : s ⊆ s + t := subset_of_le <| le_add_right s t + +lemma subset_add_right {s t : Multiset α} : s ⊆ t + s := subset_of_le <| le_add_left s t + theorem le_iff_exists_add {s t : Multiset α} : s ≤ t ↔ ∃ u, t = s + u := ⟨fun h => leInductionOn h fun s => From c4b22ac853b671670de4eb12d2c36a3692d25673 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 2 Oct 2024 11:38:30 +0000 Subject: [PATCH 116/174] feat: golf construction of stereographic projection (#17350) The calculations appearing in the construction of the stereographic projection were one of the motivating examples for the new `match_scalars` tactic. This PR rewrites those calculations to use the tactic. --- .../Geometry/Manifold/Instances/Sphere.lean | 82 +++++++------------ 1 file changed, 31 insertions(+), 51 deletions(-) diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 6b2de5ca878f6..14fe8b165b42a 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -11,6 +11,7 @@ import Mathlib.Analysis.InnerProductSpace.PiL2 import Mathlib.Geometry.Manifold.Algebra.LieGroup import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.MFDeriv.Basic +import Mathlib.Tactic.Module /-! # Manifold structure on the sphere @@ -161,7 +162,7 @@ theorem contDiff_stereoInvFunAux : ContDiff ℝ ⊤ (stereoInvFunAux v) := by have h₁ : ContDiff ℝ ⊤ fun w : E => (‖w‖ ^ 2 + 4)⁻¹ := by refine (h₀.add contDiff_const).inv ?_ intro x - nlinarith + positivity have h₂ : ContDiff ℝ ⊤ fun w => (4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v := by refine (contDiff_const.smul contDiff_id).add ?_ exact (h₀.sub contDiff_const).smul contDiff_const @@ -184,9 +185,9 @@ theorem stereoInvFun_ne_north_pole (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : rw [← inner_lt_one_iff_real_of_norm_one _ hv] · have hw : ⟪v, w⟫_ℝ = 0 := Submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2 have hw' : (‖(w : E)‖ ^ 2 + 4)⁻¹ * (‖(w : E)‖ ^ 2 - 4) < 1 := by - refine (inv_mul_lt_iff' ?_).mpr ?_ - · nlinarith - linarith + rw [inv_mul_lt_iff'] + · linarith + positivity simpa [real_inner_comm, inner_add_right, inner_smul_right, real_inner_self_eq_norm_mul_norm, hw, hv] using hw' · simpa using stereoInvFunAux_mem hv w.2 @@ -195,6 +196,7 @@ theorem continuous_stereoInvFun (hv : ‖v‖ = 1) : Continuous (stereoInvFun hv continuous_induced_rng.2 (contDiff_stereoInvFunAux.continuous.comp continuous_subtype_val) open scoped InnerProductSpace in +attribute [-simp] AddSubgroupClass.coe_norm Submodule.coe_norm in theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) ≠ v) : stereoInvFun hv (stereoToFun v x) = x := by ext @@ -212,57 +214,35 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) · simp [← split] · simp [norm_smul, hv, ← sq, sq_abs] · exact sq _ - -- two facts which will be helpful for clearing denominators in the main calculation - have ha : 1 - a ≠ 0 := by + -- a fact which will be helpful for clearing denominators in the main calculation + have ha : 0 < 1 - a := by have : a < 1 := (inner_lt_one_iff_real_of_norm_one hv (by simp)).mpr hx.symm linarith - -- the core of the problem is these two algebraic identities: - have h₁ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * 4 * (2 / (1 - a)) = 1 := by - -- TODO(#15486): used to be `field_simp`, but was really slow - -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. - simp (disch := field_simp_discharge) only [AddSubgroupClass.coe_norm, div_mul_eq_mul_div, - div_add', inv_div, mul_div_assoc', div_div, div_eq_iff, one_mul] - simp only [Submodule.coe_norm] at *; nlinarith only [pythag] - have h₂ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 - 4) = a := by - -- TODO(#15486): used to be `field_simp`, but was really slow - -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. - simp (disch := field_simp_discharge) only [AddSubgroupClass.coe_norm, div_mul_eq_mul_div, - div_add', inv_div, div_sub', mul_div_assoc', div_div, div_eq_iff] - transitivity (1 - a) ^ 2 * (a * (2 ^ 2 * ‖y‖ ^ 2 + 4 * (1 - a) ^ 2)) - · congr - simp only [Submodule.coe_norm] at * - nlinarith only [pythag] - ring! - convert - congr_arg₂ Add.add (congr_arg (fun t => t • (y : E)) h₁) (congr_arg (fun t => t • v) h₂) using 1 - · simp only [innerSL_apply, norm_smul, norm_div, RCLike.norm_ofNat, Real.norm_eq_abs, - AddSubgroupClass.coe_norm, mul_pow, div_pow, sq_abs, SetLike.val_smul, mul_smul, a] - -- Porting note: used to be simp only [split, add_comm] but get maxRec errors - rw [split, add_comm] - ac_rfl - -- Porting note: this branch did not exit in ml3 - · rw [split, add_comm] - congr - dsimp - rw [one_smul] + rw [split, Submodule.coe_smul_of_tower] + simp only [norm_smul, norm_div, RCLike.norm_ofNat, Real.norm_eq_abs, abs_of_nonneg ha.le] + match_scalars + · field_simp + linear_combination 4 * (1 - a) * pythag + · field_simp + linear_combination 4 * (a - 1) ^ 3 * pythag theorem stereo_right_inv (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : stereoToFun v (stereoInvFun hv w) = w := by - have : 2 / (1 - (‖(w : E)‖ ^ 2 + 4)⁻¹ * (‖(w : E)‖ ^ 2 - 4)) * (‖(w : E)‖ ^ 2 + 4)⁻¹ * 4 = 1 := by - -- TODO(#15486): used to be `field_simp`, but was really slow - -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. - simp (disch := field_simp_discharge) only [inv_eq_one_div, div_mul_eq_mul_div, one_mul, - sub_div', add_sub_sub_cancel, div_div_eq_mul_div, mul_div_assoc', mul_one, div_div, - div_eq_iff] - ring - convert congr_arg (· • w) this - · have h₁ : orthogonalProjection (ℝ ∙ v)ᗮ v = 0 := - orthogonalProjection_orthogonalComplement_singleton_eq_zero v - -- Porting note: was innerSL _ and now just inner - have h₃ : inner v w = (0 : ℝ) := Submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2 - -- Porting note: was innerSL _ and now just inner - have h₄ : inner v v = (1 : ℝ) := by simp [real_inner_self_eq_norm_mul_norm, hv] - simp [h₁, h₃, h₄, ContinuousLinearMap.map_add, ContinuousLinearMap.map_smul, mul_smul] - · simp + simp only [stereoToFun, stereoInvFun, stereoInvFunAux, smul_add, map_add, map_smul, innerSL_apply, + orthogonalProjection_mem_subspace_eq_self] + have h₁ : orthogonalProjection (ℝ ∙ v)ᗮ v = 0 := + orthogonalProjection_orthogonalComplement_singleton_eq_zero v + -- Porting note: was innerSL _ and now just inner + have h₂ : inner v w = (0 : ℝ) := Submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2 + -- Porting note: was innerSL _ and now just inner + have h₃ : inner v v = (1 : ℝ) := by simp [real_inner_self_eq_norm_mul_norm, hv] + rw [h₁, h₂, h₃] + match_scalars + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [add_div', add_sub_sub_cancel, div_div, + div_div_eq_mul_div, div_eq_iff, div_mul_eq_mul_div, inv_eq_one_div, + mul_div_assoc', mul_one, mul_zero, one_mul, smul_eq_mul, sub_div', zero_add, zero_div, zero_mul] + ring /-- Stereographic projection from the unit sphere in `E`, centred at a unit vector `v` in `E`; this is the version as a partial homeomorphism. -/ From 4236727d21dcfef1c709eb7801331d10755e06bc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 12:24:27 +0000 Subject: [PATCH 117/174] =?UTF-8?q?feat:=20a=20ball=20in=20`=E2=84=9D`=20i?= =?UTF-8?q?s=20a=20segment=20(#17282)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi --- Mathlib/Analysis/Convex/Topology.lean | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index 312786bb4c364..c6d92df1bdcf2 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -26,9 +26,20 @@ open Metric Bornology Set Pointwise Convex variable {ι 𝕜 E : Type*} -theorem Real.convex_iff_isPreconnected {s : Set ℝ} : Convex ℝ s ↔ IsPreconnected s := +namespace Real +variable {s : Set ℝ} {r ε : ℝ} + +lemma closedBall_eq_segment (hε : 0 ≤ ε) : closedBall r ε = segment ℝ (r - ε) (r + ε) := by + rw [closedBall_eq_Icc, segment_eq_Icc ((sub_le_self _ hε).trans <| le_add_of_nonneg_right hε)] + +lemma ball_eq_openSegment (hε : 0 < ε) : ball r ε = openSegment ℝ (r - ε) (r + ε) := by + rw [ball_eq_Ioo, openSegment_eq_Ioo ((sub_lt_self _ hε).trans <| lt_add_of_pos_right _ hε)] + +theorem convex_iff_isPreconnected : Convex ℝ s ↔ IsPreconnected s := convex_iff_ordConnected.trans isPreconnected_iff_ordConnected.symm +end Real + alias ⟨_, IsPreconnected.convex⟩ := Real.convex_iff_isPreconnected /-! ### Standard simplex -/ From 66e06b65522f07278db73e0947d72d8ae31ae1f8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 12:51:05 +0000 Subject: [PATCH 118/174] chore(SetLike): Let `gcongr` know about `coe_subset_coe` (#17290) From LeanCamCombi --- Mathlib/Data/SetLike/Basic.lean | 12 +++++------- 1 file changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index 23810a8d5caff..9613ca363c6b8 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -191,17 +191,15 @@ instance (priority := 100) instPartialOrder : PartialOrder A := theorem le_def {S T : A} : S ≤ T ↔ ∀ ⦃x : B⦄, x ∈ S → x ∈ T := Iff.rfl -@[simp, norm_cast] -theorem coe_subset_coe {S T : A} : (S : Set B) ⊆ T ↔ S ≤ T := - Iff.rfl +@[simp, norm_cast] lemma coe_subset_coe {S T : A} : (S : Set B) ⊆ T ↔ S ≤ T := .rfl +@[simp, norm_cast] lemma coe_ssubset_coe {S T : A} : (S : Set B) ⊂ T ↔ S < T := .rfl + +@[gcongr] protected alias ⟨_, GCongr.coe_subset_coe⟩ := coe_subset_coe +@[gcongr] protected alias ⟨_, GCongr.coe_ssubset_coe⟩ := coe_ssubset_coe @[mono] theorem coe_mono : Monotone (SetLike.coe : A → Set B) := fun _ _ => coe_subset_coe.mpr -@[simp, norm_cast] -theorem coe_ssubset_coe {S T : A} : (S : Set B) ⊂ T ↔ S < T := - Iff.rfl - @[mono] theorem coe_strictMono : StrictMono (SetLike.coe : A → Set B) := fun _ _ => coe_ssubset_coe.mpr From 8a443b7ad3a6c846913bb0363e9f11e82e29fce1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 14:32:01 +0000 Subject: [PATCH 119/174] =?UTF-8?q?feat:=20`Finset.sup=20s=20f=20=3D=200?= =?UTF-8?q?=20=E2=86=94=20=E2=88=80=20i=20=E2=88=88=20s,=20f=20i=20=3D=200?= =?UTF-8?q?`=20(#17078)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From ForbiddenMatrix --- Mathlib.lean | 1 + .../Algebra/Order/Monoid/Canonical/Basic.lean | 19 +++++++++++++++++++ 2 files changed, 20 insertions(+) create mode 100644 Mathlib/Algebra/Order/Monoid/Canonical/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2851c53609ce6..9c18b563e4980 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -629,6 +629,7 @@ import Mathlib.Algebra.Order.Module.Pointwise import Mathlib.Algebra.Order.Module.Rat import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Algebra.Order.Monoid.Basic +import Mathlib.Algebra.Order.Monoid.Canonical.Basic import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Algebra.Order.Monoid.NatCast diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Basic.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Basic.lean new file mode 100644 index 0000000000000..6005ed71ee499 --- /dev/null +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Basic.lean @@ -0,0 +1,19 @@ +/- +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.Order.Monoid.Canonical.Defs +import Mathlib.Data.Finset.Lattice + +/-! +# Extra lemmas about canonically ordered monoids +-/ + +namespace Finset +variable {ι α : Type*} [CanonicallyLinearOrderedAddCommMonoid α] {s : Finset ι} {f : ι → α} + +@[simp] lemma sup_eq_zero : s.sup f = 0 ↔ ∀ i ∈ s, f i = 0 := by simp [← bot_eq_zero'] +@[simp] lemma sup'_eq_zero (hs) : s.sup' hs f = 0 ↔ ∀ i ∈ s, f i = 0 := by simp [sup'_eq_sup] + +end Finset From 3b0dde03a4440144a4e587a2b0c514534d04d0f8 Mon Sep 17 00:00:00 2001 From: Etienne Date: Wed, 2 Oct 2024 14:32:03 +0000 Subject: [PATCH 120/174] feat: derivative of absolute value (#17149) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Compute the derivative of the absolute value: `deriv (|·|) x = SignType.sign x`. Co-authored-by: Etienne <66847262+EtienneC30@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Analysis/Calculus/Deriv/Abs.lean | 200 +++++++++++++++++++++++ Mathlib/Analysis/Calculus/Deriv/Add.lean | 10 -- 3 files changed, 201 insertions(+), 10 deletions(-) create mode 100644 Mathlib/Analysis/Calculus/Deriv/Abs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9c18b563e4980..5ec2f99349687 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -987,6 +987,7 @@ import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.ContDiff.FiniteDimension import Mathlib.Analysis.Calculus.ContDiff.RCLike import Mathlib.Analysis.Calculus.Darboux +import Mathlib.Analysis.Calculus.Deriv.Abs import Mathlib.Analysis.Calculus.Deriv.Add import Mathlib.Analysis.Calculus.Deriv.AffineMap import Mathlib.Analysis.Calculus.Deriv.Basic diff --git a/Mathlib/Analysis/Calculus/Deriv/Abs.lean b/Mathlib/Analysis/Calculus/Deriv/Abs.lean new file mode 100644 index 0000000000000..55365ae78213a --- /dev/null +++ b/Mathlib/Analysis/Calculus/Deriv/Abs.lean @@ -0,0 +1,200 @@ +/- +Copyright (c) 2024 Etienne Marion. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Etienne Marion +-/ +import Mathlib.Analysis.Calculus.Deriv.Add +import Mathlib.Analysis.InnerProductSpace.Calculus + +/-! +# Derivative of the absolute value + +This file compiles basic derivability properties of the absolute value, and is largely inspired +from `Mathlib.Analysis.InnerProductSpace.Calculus`, which is the analoguous file for norms derived +from an inner product space. + +## Tags + +absolute value, derivative +-/ + +open Filter Real Set + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] +variable {n : ℕ∞} {f g : E → ℝ} {f' : E →L[ℝ] ℝ} {s : Set E} {x : E} + +theorem contDiffAt_abs {x : ℝ} (hx : x ≠ 0) : ContDiffAt ℝ n (|·|) x := contDiffAt_norm ℝ hx + +theorem ContDiffAt.abs (hf : ContDiffAt ℝ n f x) (h₀ : f x ≠ 0) : + ContDiffAt ℝ n (fun x ↦ |f x|) x := hf.norm ℝ h₀ + +theorem contDiffWithinAt_abs {x : ℝ} (hx : x ≠ 0) (s : Set ℝ) : + ContDiffWithinAt ℝ n (|·|) s x := (contDiffAt_abs hx).contDiffWithinAt + +theorem ContDiffWithinAt.abs (hf : ContDiffWithinAt ℝ n f s x) (h₀ : f x ≠ 0) : + ContDiffWithinAt ℝ n (fun y ↦ |f y|) s x := + (contDiffAt_abs h₀).comp_contDiffWithinAt x hf + +theorem contDiffOn_abs {s : Set ℝ} (hs : ∀ x ∈ s, x ≠ 0) : + ContDiffOn ℝ n (|·|) s := fun x hx ↦ contDiffWithinAt_abs (hs x hx) s + +theorem ContDiffOn.abs (hf : ContDiffOn ℝ n f s) (h₀ : ∀ x ∈ s, f x ≠ 0) : + ContDiffOn ℝ n (fun y ↦ |f y|) s := fun x hx ↦ (hf x hx).abs (h₀ x hx) + +theorem ContDiff.abs (hf : ContDiff ℝ n f) (h₀ : ∀ x, f x ≠ 0) : ContDiff ℝ n fun y ↦ |f y| := + contDiff_iff_contDiffAt.2 fun x ↦ hf.contDiffAt.abs (h₀ x) + +theorem hasStrictDerivAt_abs_neg {x : ℝ} (hx : x < 0) : + HasStrictDerivAt (|·|) (-1) x := + (hasStrictDerivAt_neg x).congr_of_eventuallyEq <| + EqOn.eventuallyEq_of_mem (fun _ hy ↦ (abs_of_neg (mem_Iio.1 hy)).symm) (Iio_mem_nhds hx) + +theorem hasDerivAt_abs_neg {x : ℝ} (hx : x < 0) : + HasDerivAt (|·|) (-1) x := (hasStrictDerivAt_abs_neg hx).hasDerivAt + +theorem hasStrictDerivAt_abs_pos {x : ℝ} (hx : 0 < x) : + HasStrictDerivAt (|·|) 1 x := + (hasStrictDerivAt_id x).congr_of_eventuallyEq <| + EqOn.eventuallyEq_of_mem (fun _ hy ↦ (abs_of_pos (mem_Iio.1 hy)).symm) (Ioi_mem_nhds hx) + +theorem hasDerivAt_abs_pos {x : ℝ} (hx : 0 < x) : + HasDerivAt (|·|) 1 x := (hasStrictDerivAt_abs_pos hx).hasDerivAt + +theorem hasStrictDerivAt_abs {x : ℝ} (hx : x ≠ 0) : + HasStrictDerivAt (|·|) (SignType.sign x : ℝ) x := by + obtain hx | hx := hx.lt_or_lt + · simpa [hx] using hasStrictDerivAt_abs_neg hx + · simpa [hx] using hasStrictDerivAt_abs_pos hx + +theorem hasDerivAt_abs {x : ℝ} (hx : x ≠ 0) : + HasDerivAt (|·|) (SignType.sign x : ℝ) x := (hasStrictDerivAt_abs hx).hasDerivAt + +theorem HasStrictFDerivAt.abs_of_neg (hf : HasStrictFDerivAt f f' x) + (h₀ : f x < 0) : HasStrictFDerivAt (fun x ↦ |f x|) (-f') x := by + convert (hasStrictDerivAt_abs_neg h₀).hasStrictFDerivAt.comp x hf using 1 + ext y + simp + +theorem HasFDerivAt.abs_of_neg (hf : HasFDerivAt f f' x) + (h₀ : f x < 0) : HasFDerivAt (fun x ↦ |f x|) (-f') x := by + convert (hasDerivAt_abs_neg h₀).hasFDerivAt.comp x hf using 1 + ext y + simp + +theorem HasStrictFDerivAt.abs_of_pos (hf : HasStrictFDerivAt f f' x) + (h₀ : 0 < f x) : HasStrictFDerivAt (fun x ↦ |f x|) f' x := by + convert (hasStrictDerivAt_abs_pos h₀).hasStrictFDerivAt.comp x hf using 1 + ext y + simp + +theorem HasFDerivAt.abs_of_pos (hf : HasFDerivAt f f' x) + (h₀ : 0 < f x) : HasFDerivAt (fun x ↦ |f x|) f' x := by + convert (hasDerivAt_abs_pos h₀).hasFDerivAt.comp x hf using 1 + ext y + simp + +theorem HasStrictFDerivAt.abs (hf : HasStrictFDerivAt f f' x) + (h₀ : f x ≠ 0) : HasStrictFDerivAt (fun x ↦ |f x|) ((SignType.sign (f x) : ℝ) • f') x := by + convert (hasStrictDerivAt_abs h₀).hasStrictFDerivAt.comp x hf using 1 + ext y + simp [mul_comm] + +theorem HasFDerivAt.abs (hf : HasFDerivAt f f' x) + (h₀ : f x ≠ 0) : HasFDerivAt (fun x ↦ |f x|) ((SignType.sign (f x) : ℝ) • f') x := by + convert (hasDerivAt_abs h₀).hasFDerivAt.comp x hf using 1 + ext y + simp [mul_comm] + +theorem hasDerivWithinAt_abs_neg (s : Set ℝ) {x : ℝ} (hx : x < 0) : + HasDerivWithinAt (|·|) (-1) s x := (hasDerivAt_abs_neg hx).hasDerivWithinAt + +theorem hasDerivWithinAt_abs_pos (s : Set ℝ) {x : ℝ} (hx : 0 < x) : + HasDerivWithinAt (|·|) 1 s x := (hasDerivAt_abs_pos hx).hasDerivWithinAt + +theorem hasDerivWithinAt_abs (s : Set ℝ) {x : ℝ} (hx : x ≠ 0) : + HasDerivWithinAt (|·|) (SignType.sign x : ℝ) s x := (hasDerivAt_abs hx).hasDerivWithinAt + +theorem HasFDerivWithinAt.abs_of_neg (hf : HasFDerivWithinAt f f' s x) + (h₀ : f x < 0) : HasFDerivWithinAt (fun x ↦ |f x|) (-f') s x := by + convert (hasDerivAt_abs_neg h₀).comp_hasFDerivWithinAt x hf using 1 + simp + +theorem HasFDerivWithinAt.abs_of_pos (hf : HasFDerivWithinAt f f' s x) + (h₀ : 0 < f x) : HasFDerivWithinAt (fun x ↦ |f x|) f' s x := by + convert (hasDerivAt_abs_pos h₀).comp_hasFDerivWithinAt x hf using 1 + simp + +theorem HasFDerivWithinAt.abs (hf : HasFDerivWithinAt f f' s x) + (h₀ : f x ≠ 0) : HasFDerivWithinAt (fun x ↦ |f x|) ((SignType.sign (f x) : ℝ) • f') s x := + (hasDerivAt_abs h₀).comp_hasFDerivWithinAt x hf + +theorem differentiableAt_abs_neg {x : ℝ} (hx : x < 0) : + DifferentiableAt ℝ (|·|) x := (hasDerivAt_abs_neg hx).differentiableAt + +theorem differentiableAt_abs_pos {x : ℝ} (hx : 0 < x) : + DifferentiableAt ℝ (|·|) x := (hasDerivAt_abs_pos hx).differentiableAt + +theorem differentiableAt_abs {x : ℝ} (hx : x ≠ 0) : + DifferentiableAt ℝ (|·|) x := (hasDerivAt_abs hx).differentiableAt + +theorem DifferentiableAt.abs_of_neg (hf : DifferentiableAt ℝ f x) (h₀ : f x < 0) : + DifferentiableAt ℝ (fun x ↦ |f x|) x := (differentiableAt_abs_neg h₀).comp x hf + +theorem DifferentiableAt.abs_of_pos (hf : DifferentiableAt ℝ f x) (h₀ : 0 < f x) : + DifferentiableAt ℝ (fun x ↦ |f x|) x := (differentiableAt_abs_pos h₀).comp x hf + +theorem DifferentiableAt.abs (hf : DifferentiableAt ℝ f x) (h₀ : f x ≠ 0) : + DifferentiableAt ℝ (fun x ↦ |f x|) x := (differentiableAt_abs h₀).comp x hf + +theorem differentiableWithinAt_abs_neg (s : Set ℝ) {x : ℝ} (hx : x < 0) : + DifferentiableWithinAt ℝ (|·|) s x := (differentiableAt_abs_neg hx).differentiableWithinAt + +theorem differentiableWithinAt_abs_pos (s : Set ℝ) {x : ℝ} (hx : 0 < x) : + DifferentiableWithinAt ℝ (|·|) s x := (differentiableAt_abs_pos hx).differentiableWithinAt + +theorem differentiableWithinAt_abs (s : Set ℝ) {x : ℝ} (hx : x ≠ 0) : + DifferentiableWithinAt ℝ (|·|) s x := (differentiableAt_abs hx).differentiableWithinAt + +theorem DifferentiableWithinAt.abs_of_neg (hf : DifferentiableWithinAt ℝ f s x) (h₀ : f x < 0) : + DifferentiableWithinAt ℝ (fun x ↦ |f x|) s x := + (differentiableAt_abs_neg h₀).comp_differentiableWithinAt x hf + +theorem DifferentiableWithinAt.abs_of_pos (hf : DifferentiableWithinAt ℝ f s x) (h₀ : 0 < f x) : + DifferentiableWithinAt ℝ (fun x ↦ |f x|) s x := + (differentiableAt_abs_pos h₀).comp_differentiableWithinAt x hf + +theorem DifferentiableWithinAt.abs (hf : DifferentiableWithinAt ℝ f s x) (h₀ : f x ≠ 0) : + DifferentiableWithinAt ℝ (fun x ↦ |f x|) s x := + (differentiableAt_abs h₀).comp_differentiableWithinAt x hf + +theorem differentiableOn_abs {s : Set ℝ} (hs : ∀ x ∈ s, x ≠ 0) : DifferentiableOn ℝ (|·|) s := + fun x hx ↦ differentiableWithinAt_abs s (hs x hx) + +theorem DifferentiableOn.abs (hf : DifferentiableOn ℝ f s) (h₀ : ∀ x ∈ s, f x ≠ 0) : + DifferentiableOn ℝ (fun x ↦ |f x|) s := + fun x hx ↦ (hf x hx).abs (h₀ x hx) + +theorem Differentiable.abs (hf : Differentiable ℝ f) (h₀ : ∀ x, f x ≠ 0) : + Differentiable ℝ (fun x ↦ |f x|) := fun x ↦ (hf x).abs (h₀ x) + +theorem not_differentiableAt_abs_zero : ¬ DifferentiableAt ℝ (abs : ℝ → ℝ) 0 := by + intro h + have h₁ : deriv abs (0 : ℝ) = 1 := + (uniqueDiffOn_Ici _ _ Set.left_mem_Ici).eq_deriv _ h.hasDerivAt.hasDerivWithinAt <| + (hasDerivWithinAt_id _ _).congr_of_mem (fun _ h ↦ abs_of_nonneg h) Set.left_mem_Ici + have h₂ : deriv abs (0 : ℝ) = -1 := + (uniqueDiffOn_Iic _ _ Set.right_mem_Iic).eq_deriv _ h.hasDerivAt.hasDerivWithinAt <| + (hasDerivWithinAt_neg _ _).congr_of_mem (fun _ h ↦ abs_of_nonpos h) Set.right_mem_Iic + linarith + +theorem deriv_abs_neg {x : ℝ} (hx : x < 0) : deriv (|·|) x = -1 := (hasDerivAt_abs_neg hx).deriv + +theorem deriv_abs_pos {x : ℝ} (hx : 0 < x) : deriv (|·|) x = 1 := (hasDerivAt_abs_pos hx).deriv + +theorem deriv_abs_zero : deriv (|·|) (0 : ℝ) = 0 := + deriv_zero_of_not_differentiableAt not_differentiableAt_abs_zero + +theorem deriv_abs (x : ℝ) : deriv (|·|) x = SignType.sign x := by + obtain rfl | hx := eq_or_ne x 0 + · simpa using deriv_abs_zero + · simpa [hx] using (hasDerivAt_abs hx).deriv diff --git a/Mathlib/Analysis/Calculus/Deriv/Add.lean b/Mathlib/Analysis/Calculus/Deriv/Add.lean index 6f51814995811..169722bc11409 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Add.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Add.lean @@ -247,16 +247,6 @@ theorem differentiable_neg : Differentiable 𝕜 (Neg.neg : 𝕜 → 𝕜) := theorem differentiableOn_neg : DifferentiableOn 𝕜 (Neg.neg : 𝕜 → 𝕜) s := DifferentiableOn.neg differentiableOn_id -theorem not_differentiableAt_abs_zero : ¬ DifferentiableAt ℝ (abs : ℝ → ℝ) 0 := by - intro h - have h₁ : deriv abs (0 : ℝ) = 1 := - (uniqueDiffOn_Ici _ _ Set.left_mem_Ici).eq_deriv _ h.hasDerivAt.hasDerivWithinAt <| - (hasDerivWithinAt_id _ _).congr_of_mem (fun _ h ↦ abs_of_nonneg h) Set.left_mem_Ici - have h₂ : deriv abs (0 : ℝ) = -1 := - (uniqueDiffOn_Iic _ _ Set.right_mem_Iic).eq_deriv _ h.hasDerivAt.hasDerivWithinAt <| - (hasDerivWithinAt_neg _ _).congr_of_mem (fun _ h ↦ abs_of_nonpos h) Set.right_mem_Iic - linarith - lemma differentiableAt_comp_neg {a : 𝕜} : DifferentiableAt 𝕜 (fun x ↦ f (-x)) a ↔ DifferentiableAt 𝕜 f (-a) := by refine ⟨fun H ↦ ?_, fun H ↦ H.comp a differentiable_neg.differentiableAt⟩ From 9aacea7f101b62b2aae697b2774edb28e2c1bdd5 Mon Sep 17 00:00:00 2001 From: JonBannon Date: Wed, 2 Oct 2024 14:32:04 +0000 Subject: [PATCH 121/174] feat(Analysis.InnerProductSpace.Symmetric) : Added several `aesop` lemmas for operations preserving `IsSymmetric` (#17255) Added `@[aesop safe apply]` to `isSymmetric_zero, isSymmetric_id, IsSymmetric.add`, as well as to the following newly added theorems: `isSymmetric.sub`, `isSymmetric.smul`, `isSymmetric.mul_of_comm`, `isSymmetric.pow`. --- .../Analysis/InnerProductSpace/Symmetric.lean | 35 +++++++++++++++++-- 1 file changed, 33 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 0dafa0de1b246..9acf7b33a0635 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -74,17 +74,48 @@ theorem IsSymmetric.apply_clm {T : E →L[𝕜] E} (hT : IsSymmetric (T : E → ⟪T x, y⟫ = ⟪x, T y⟫ := hT x y -theorem isSymmetric_zero : (0 : E →ₗ[𝕜] E).IsSymmetric := fun x y => +@[simp] +protected theorem IsSymmetric.zero : (0 : E →ₗ[𝕜] E).IsSymmetric := fun x y => (inner_zero_right x : ⟪x, 0⟫ = 0).symm ▸ (inner_zero_left y : ⟪0, y⟫ = 0) -theorem isSymmetric_id : (LinearMap.id : E →ₗ[𝕜] E).IsSymmetric := fun _ _ => rfl +@[deprecated (since := "2024-09-30")] alias isSymmetric_zero := IsSymmetric.zero + +@[simp] +protected theorem IsSymmetric.id : (LinearMap.id : E →ₗ[𝕜] E).IsSymmetric := fun _ _ => rfl +@[deprecated (since := "2024-09-30")] alias isSymmetric_id := IsSymmetric.id + +@[aesop safe apply] theorem IsSymmetric.add {T S : E →ₗ[𝕜] E} (hT : T.IsSymmetric) (hS : S.IsSymmetric) : (T + S).IsSymmetric := by intro x y rw [LinearMap.add_apply, inner_add_left, hT x y, hS x y, ← inner_add_right] rfl +@[aesop safe apply] +theorem IsSymmetric.sub {T S : E →ₗ[𝕜] E} (hT : T.IsSymmetric) (hS : S.IsSymmetric) : + (T - S).IsSymmetric := by + intro x y + rw [LinearMap.sub_apply, inner_sub_left, hT x y, hS x y, ← inner_sub_right] + rfl + +@[aesop safe apply] +theorem IsSymmetric.smul {c : 𝕜} (hc : conj c = c) {T : E →ₗ[𝕜] E} (hT : T.IsSymmetric) : + c • T |>.IsSymmetric := by + intro x y + simp only [smul_apply, inner_smul_left, hc, hT x y, inner_smul_right] + +@[aesop 30% apply] +lemma IsSymmetric.mul_of_commute {S T : E →ₗ[𝕜] E} (hS : S.IsSymmetric) (hT : T.IsSymmetric) + (hST : Commute S T) : (S * T).IsSymmetric := + fun _ _ ↦ by rw [mul_apply, hS, hT, hST, mul_apply] + +@[aesop safe apply] +lemma IsSymmetric.pow {T : E →ₗ[𝕜] E} (hT : T.IsSymmetric) (n : ℕ) : (T ^ n).IsSymmetric := by + refine Nat.le_induction (by simp [one_eq_id]) (fun k _ ih ↦ ?_) n n.zero_le + rw [iterate_succ, ← mul_eq_comp] + exact ih.mul_of_commute hT <| .pow_left rfl k + /-- For a symmetric operator `T`, the function `fun x ↦ ⟪T x, x⟫` is real-valued. -/ @[simp] theorem IsSymmetric.coe_reApplyInnerSelf_apply {T : E →L[𝕜] E} (hT : IsSymmetric (T : E →ₗ[𝕜] E)) From c472200bb60e0adf2bf0cea4b0d04985511b276d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 14:57:40 +0000 Subject: [PATCH 122/174] feat: `@[deprecated]` code snippet (#17357) We already had a `@[deprecated] alias` code snippet. This one is useful when deprecating a lemma in place. --- ...ecated-alias.code-snippets => deprecated.code-snippets} | 7 +++++++ 1 file changed, 7 insertions(+) rename .vscode/{deprecated-alias.code-snippets => deprecated.code-snippets} (53%) diff --git a/.vscode/deprecated-alias.code-snippets b/.vscode/deprecated.code-snippets similarity index 53% rename from .vscode/deprecated-alias.code-snippets rename to .vscode/deprecated.code-snippets index c760b46045763..8be8e537cd5f0 100644 --- a/.vscode/deprecated-alias.code-snippets +++ b/.vscode/deprecated.code-snippets @@ -1,4 +1,11 @@ { + "Deprecation for mathlib": { + "scope": "lean4", + "prefix": "deprecated", + "body": [ + "@[deprecated $1 (since := \"${CURRENT_YEAR}-${CURRENT_MONTH}-${CURRENT_DATE}\")]" + ] + }, "Deprecated alias for mathlib": { "scope": "lean4", "prefix": "deprecated alias", From 4660dcd5b4234594c7ebf2bcfb56f19c6bf156f3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 16:36:02 +0000 Subject: [PATCH 123/174] feat: there are finitely many lists of a fixed length (#17288) From LeanCamCombi --- Mathlib/Data/Set/Finite.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index 2cc22fdd69e8c..dcd9b53e9dcc2 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -1528,6 +1528,7 @@ protected theorem bddBelow [SemilatticeInf α] [Nonempty α] (s : Finset α) : B end Finset +section LinearOrder variable [LinearOrder α] {s : Set α} /-- If a linear order does not contain any triple of elements `x < y < z`, then this type @@ -1568,4 +1569,19 @@ theorem DirectedOn.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} {f haveI := hn.coe_sort simpa using (directed_comp.2 hc.directed_val).exists_mem_subset_of_finset_subset_biUnion hs +end LinearOrder + +namespace List +variable (α) [Finite α] (n : ℕ) + +lemma finite_length_eq : {l : List α | l.length = n}.Finite := Vector.finite + +lemma finite_length_lt : {l : List α | l.length < n}.Finite := by + convert (Finset.range n).finite_toSet.biUnion fun i _ ↦ finite_length_eq α i; ext; simp + +lemma finite_length_le : {l : List α | l.length ≤ n}.Finite := by + simpa [Nat.lt_succ_iff] using finite_length_lt α (n + 1) + +end List + set_option linter.style.longFile 1700 From 343c459e02fb933bc42a6aff3a3554e6dad168cd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 2 Oct 2024 17:07:53 +0000 Subject: [PATCH 124/174] feat(Analysis.Normed.Ring.IsPowMulFaithful): prove eq_of_powMul_faithful (#15445) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We add `AlgebraNorm` and `MulAlgebraNorm`. We prove Bosch-Güntzer-Remmert, Proposition 3.1.5/1: if `R` is a normed commutative ring and `f₁` and `f₂` are two power-multiplicative `R`-algebra norms on `S`, then if `f₁` and `f₂` are equivalent on every subring `R[y]` for `y : S`, it follows that `f₁ = f₂`. --- Mathlib.lean | 2 + Mathlib/Analysis/Normed/Algebra/Norm.lean | 202 ++++++++++++++++++ Mathlib/Analysis/Normed/Field/Basic.lean | 12 ++ .../Normed/Ring/IsPowMulFaithful.lean | 96 +++++++++ Mathlib/Analysis/Normed/Ring/Seminorm.lean | 53 +++++ Mathlib/Data/Real/Basic.lean | 7 + 6 files changed, 372 insertions(+) create mode 100644 Mathlib/Analysis/Normed/Algebra/Norm.lean create mode 100644 Mathlib/Analysis/Normed/Ring/IsPowMulFaithful.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5ec2f99349687..0d6d6971bf5da 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1201,6 +1201,7 @@ import Mathlib.Analysis.Normed.Affine.MazurUlam import Mathlib.Analysis.Normed.Algebra.Basic import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Normed.Algebra.MatrixExponential +import Mathlib.Analysis.Normed.Algebra.Norm import Mathlib.Analysis.Normed.Algebra.QuaternionExponential import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.Normed.Algebra.TrivSqZeroExt @@ -1262,6 +1263,7 @@ import Mathlib.Analysis.Normed.Operator.WeakOperatorTopology import Mathlib.Analysis.Normed.Order.Basic import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.Normed.Order.UpperLower +import Mathlib.Analysis.Normed.Ring.IsPowMulFaithful import Mathlib.Analysis.Normed.Ring.Seminorm import Mathlib.Analysis.Normed.Ring.SeminormFromBounded import Mathlib.Analysis.Normed.Ring.SeminormFromConst diff --git a/Mathlib/Analysis/Normed/Algebra/Norm.lean b/Mathlib/Analysis/Normed/Algebra/Norm.lean new file mode 100644 index 0000000000000..eb815ac91298e --- /dev/null +++ b/Mathlib/Analysis/Normed/Algebra/Norm.lean @@ -0,0 +1,202 @@ +/- +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.Analysis.Normed.Ring.Seminorm +import Mathlib.Analysis.Seminorm + +/-! +# Algebra norms + +We define algebra norms and multiplicative algebra norms. + +## Main Definitions +* `AlgebraNorm` : an algebra norm on an `R`-algebra `S` is a ring norm on `S` compatible with + the action of `R`. +* `MulAlgebraNorm` : a multiplicative algebra norm on an `R`-algebra `S` is a multiplicative + ring norm on `S` compatible with the action of `R`. + +## Tags + +norm, algebra norm +-/ + +/-- An algebra norm on an `R`-algebra `S` is a ring norm on `S` compatible with the +action of `R`. -/ +structure AlgebraNorm (R : Type*) [SeminormedCommRing R] (S : Type*) [Ring S] [Algebra R S] extends + RingNorm S, Seminorm R S + +attribute [nolint docBlame] AlgebraNorm.toSeminorm AlgebraNorm.toRingNorm + +instance (K : Type*) [NormedField K] : Inhabited (AlgebraNorm K K) := + ⟨{ toFun := norm + map_zero' := norm_zero + add_le' := norm_add_le + neg' := norm_neg + smul' := norm_mul + mul_le' := norm_mul_le + eq_zero_of_map_eq_zero' := fun _ => norm_eq_zero.mp }⟩ + +/-- `AlgebraNormClass F R S` states that `F` is a type of `R`-algebra norms on the ring `S`. +You should extend this class when you extend `AlgebraNorm`. -/ +class AlgebraNormClass (F : Type*) (R : outParam <| Type*) [SeminormedCommRing R] + (S : outParam <| Type*) [Ring S] [Algebra R S] [FunLike F S ℝ] extends RingNormClass F S ℝ, + SeminormClass F R S : Prop + +namespace AlgebraNorm + +variable {R : Type*} [SeminormedCommRing R] {S : Type*} [Ring S] [Algebra R S] {f : AlgebraNorm R S} + +/-- The ring seminorm underlying an algebra norm. -/ +def toRingSeminorm' (f : AlgebraNorm R S) : RingSeminorm S := + f.toRingNorm.toRingSeminorm + +instance : FunLike (AlgebraNorm R S) S ℝ where + coe f := f.toFun + coe_injective' f f' h := by + simp only [AddGroupSeminorm.toFun_eq_coe, RingSeminorm.toFun_eq_coe] at h + cases f; cases f'; congr; + simp only at h + ext s + erw [h] + rfl + +instance algebraNormClass : AlgebraNormClass (AlgebraNorm R S) R S where + map_zero f := f.map_zero' + map_add_le_add f := f.add_le' + map_mul_le_mul f := f.mul_le' + map_neg_eq_map f := f.neg' + 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] +theorem ext {p q : AlgebraNorm R S} : (∀ x, p x = q x) → p = q := + DFunLike.ext p q + +/-- An `R`-algebra norm such that `f 1 = 1` extends the norm on `R`. -/ +theorem extends_norm' (hf1 : f 1 = 1) (a : R) : f (a • (1 : S)) = ‖a‖ := by + rw [← mul_one ‖a‖, ← hf1]; exact f.smul' _ _ + +/-- An `R`-algebra norm such that `f 1 = 1` extends the norm on `R`. -/ +theorem extends_norm (hf1 : f 1 = 1) (a : R) : f (algebraMap R S a) = ‖a‖ := by + rw [Algebra.algebraMap_eq_smul_one]; exact extends_norm' hf1 _ + +/-- The restriction of an algebra norm to a subalgebra. -/ +def restriction (A : Subalgebra R S) (f : AlgebraNorm R S) : AlgebraNorm R A where + toFun := fun x : A => f x.val + map_zero' := map_zero f + add_le' x y := map_add_le_add _ _ _ + neg' x := map_neg_eq_map _ _ + mul_le' x y := map_mul_le_mul _ _ _ + eq_zero_of_map_eq_zero' x hx := by + rw [← ZeroMemClass.coe_eq_zero]; exact eq_zero_of_map_eq_zero f hx + smul' r x := map_smul_eq_mul _ _ _ + +/-- The restriction of an algebra norm in a scalar tower. -/ +def isScalarTower_restriction {A : Type*} [CommRing A] [Algebra R A] [Algebra A S] + [IsScalarTower R A S] (hinj : Function.Injective (algebraMap A S)) (f : AlgebraNorm R S) : + AlgebraNorm R A where + toFun := fun x : A => f (algebraMap A S x) + map_zero' := by simp only [map_zero] + add_le' x y := by simp only [map_add, map_add_le_add] + neg' x := by simp only [map_neg, map_neg_eq_map] + mul_le' x y := by simp only [map_mul, map_mul_le_mul] + eq_zero_of_map_eq_zero' x hx := by + rw [← map_eq_zero_iff (algebraMap A S) hinj] + exact eq_zero_of_map_eq_zero f hx + smul' r x := by + simp only [Algebra.smul_def, map_mul, ← IsScalarTower.algebraMap_apply] + simp only [← smul_eq_mul, algebraMap_smul, map_smul_eq_mul] + +end AlgebraNorm + +/-- A multiplicative algebra norm on an `R`-algebra norm `S` is a multiplicative ring norm on `S` + compatible with the action of `R`. -/ +structure MulAlgebraNorm (R : Type*) [SeminormedCommRing R] (S : Type*) [Ring S] [Algebra R S] + extends MulRingNorm S, Seminorm R S + +attribute [nolint docBlame] MulAlgebraNorm.toSeminorm MulAlgebraNorm.toMulRingNorm + +instance (K : Type*) [NormedField K] : Inhabited (MulAlgebraNorm K K) := + ⟨{ toFun := norm + map_zero' := norm_zero + add_le' := norm_add_le + neg' := norm_neg + smul' := norm_mul + map_one' := norm_one + map_mul' := norm_mul + eq_zero_of_map_eq_zero' := fun _ => norm_eq_zero.mp }⟩ + +/-- `MulAlgebraNormClass F R S` states that `F` is a type of multiplicative `R`-algebra norms on +the ring `S`. You should extend this class when you extend `MulAlgebraNorm`. -/ +class MulAlgebraNormClass (F : Type*) (R : outParam <| Type*) [SeminormedCommRing R] + (S : outParam <| Type*) [Ring S] [Algebra R S] [FunLike F S ℝ] extends MulRingNormClass F S ℝ, + SeminormClass F R S : Prop + +namespace MulAlgebraNorm + +variable {R S : outParam <| Type*} [SeminormedCommRing R] [Ring S] [Algebra R S] + {f : AlgebraNorm R S} + +instance : FunLike (MulAlgebraNorm R S) S ℝ where + coe f := f.toFun + coe_injective' f f' h:= by + simp only [AddGroupSeminorm.toFun_eq_coe, MulRingSeminorm.toFun_eq_coe, DFunLike.coe_fn_eq] at h + obtain ⟨⟨_, _⟩, _⟩ := f; obtain ⟨⟨_, _⟩, _⟩ := f'; congr; + +instance mulAlgebraNormClass : MulAlgebraNormClass (MulAlgebraNorm R S) R S where + map_zero f := f.map_zero' + map_add_le_add f := f.add_le' + map_one f := f.map_one' + map_mul f := f.map_mul' + map_neg_eq_map f := f.neg' + 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] +theorem ext {p q : MulAlgebraNorm R S} : (∀ x, p x = q x) → p = q := + DFunLike.ext p q + +/-- A multiplicative `R`-algebra norm extends the norm on `R`. -/ +theorem extends_norm' (f : MulAlgebraNorm R S) (a : R) : f (a • (1 : S)) = ‖a‖ := by + rw [← mul_one ‖a‖, ← f.map_one', ← f.smul']; rfl + +/-- A multiplicative `R`-algebra norm extends the norm on `R`. -/ +theorem extends_norm (f : MulAlgebraNorm R S) (a : R) : f (algebraMap R S a) = ‖a‖ := by + rw [Algebra.algebraMap_eq_smul_one]; exact extends_norm' _ _ + +end MulAlgebraNorm + +namespace MulRingNorm + +variable {R : Type*} [NonAssocRing R] + +/-- The ring norm underlying a multiplicative ring norm. -/ +def toRingNorm (f : MulRingNorm R) : RingNorm R where + toFun := f + map_zero' := f.map_zero' + add_le' := f.add_le' + neg' := f.neg' + mul_le' x y := le_of_eq (f.map_mul' x y) + eq_zero_of_map_eq_zero' := f.eq_zero_of_map_eq_zero' + +/-- A multiplicative ring norm is power-multiplicative. -/ +theorem isPowMul {A : Type*} [Ring A] (f : MulRingNorm A) : IsPowMul f := fun x n hn => by + cases n + · exfalso; linarith + · rw [map_pow] + +end MulRingNorm diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index a8762297c678a..cb044efa45b52 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -443,6 +443,12 @@ 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 +/-- 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 β] + (f : α →+* β) : Prop := + ∃ C : ℝ, 0 < C ∧ ∀ x : α, norm (f x) ≤ C * norm x + end SeminormedRing section NonUnitalNormedRing @@ -586,6 +592,12 @@ instance MulOpposite.instNormedCommRing : NormedCommRing αᵐᵒᵖ where __ := instNormedRing __ := instSeminormedCommRing +/-- The restriction of a power-multiplicative function to a subalgebra is power-multiplicative. -/ +theorem IsPowMul.restriction {R S : Type*} [NormedCommRing R] [CommRing S] [Algebra R S] + (A : Subalgebra R S) {f : S → ℝ} (hf_pm : IsPowMul f) : + IsPowMul fun x : A => f x.val := fun x n hn => by + simpa [SubsemiringClass.coe_pow] using hf_pm (↑x) hn + end NormedCommRing section NormedDivisionRing diff --git a/Mathlib/Analysis/Normed/Ring/IsPowMulFaithful.lean b/Mathlib/Analysis/Normed/Ring/IsPowMulFaithful.lean new file mode 100644 index 0000000000000..615d91284df1a --- /dev/null +++ b/Mathlib/Analysis/Normed/Ring/IsPowMulFaithful.lean @@ -0,0 +1,96 @@ +/- +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.Analysis.Normed.Algebra.Norm +import Mathlib.Analysis.SpecialFunctions.Pow.Continuity + +/-! +# Equivalent power-multiplicative norms + +In this file, we prove [BGR, Proposition 3.1.5/1][bosch-guntzer-remmert]: if `R` is a normed +commutative ring and `f₁` and `f₂` are two power-multiplicative `R`-algebra norms on `S`, then if +`f₁` and `f₂` are equivalent on every subring `R[y]` for `y : S`, it follows that `f₁ = f₂`. + +## Main Results +* `eq_of_powMul_faithful` : the proof of [BGR, Proposition 3.1.5/1][bosch-guntzer-remmert]. + +## References +* [S. Bosch, U. Güntzer, R. Remmert, *Non-Archimedean Analysis*][bosch-guntzer-remmert] + +## Tags + +norm, equivalent, power-multiplicative +-/ + +open Filter Real +open scoped Topology + +/-- If `f : α →+* β` is bounded with respect to a ring seminorm `nα` on `α` and a + power-multiplicative function `nβ : β → ℝ`, then `∀ x : α, nβ (f x) ≤ nα x`. -/ +theorem contraction_of_isPowMul_of_boundedWrt {F : Type*} {α : outParam (Type*)} [Ring α] + [FunLike F α ℝ] [RingSeminormClass F α ℝ] {β : Type*} [Ring β] (nα : F) {nβ : β → ℝ} + (hβ : IsPowMul nβ) {f : α →+* β} (hf : f.IsBoundedWrt nα nβ) (x : α) : nβ (f x) ≤ nα x := by + obtain ⟨C, hC0, hC⟩ := hf + have hlim : Tendsto (fun n : ℕ => C ^ (1 / (n : ℝ)) * nα x) atTop (𝓝 (nα x)) := by + nth_rewrite 2 [← one_mul (nα x)] + exact ((rpow_zero C ▸ ContinuousAt.tendsto (continuousAt_const_rpow (ne_of_gt hC0))).comp + (tendsto_const_div_atTop_nhds_zero_nat 1)).mul tendsto_const_nhds + apply ge_of_tendsto hlim + simp only [eventually_atTop, ge_iff_le] + use 1 + intro n hn + have h : (C ^ (1 / n : ℝ)) ^ n = C := by + have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.mpr (ne_of_gt hn) + rw [← rpow_natCast, ← rpow_mul (le_of_lt hC0), one_div, inv_mul_cancel₀ hn0, rpow_one] + apply le_of_pow_le_pow_left (ne_of_gt hn) + (mul_nonneg (rpow_nonneg (le_of_lt hC0) _) (apply_nonneg _ _)) + · rw [mul_pow, h, ← hβ _ hn, ← RingHom.map_pow] + apply le_trans (hC (x ^ n)) + rw [mul_le_mul_left hC0] + exact map_pow_le_pow _ _ (Nat.one_le_iff_ne_zero.mp hn) + +/-- Given a bounded `f : α →+* β` between seminormed rings, is the seminorm on `β` is + power-multiplicative, then `f` is a contraction. -/ +theorem contraction_of_isPowMul {α β : Type*} [SeminormedRing α] [SeminormedRing β] + (hβ : IsPowMul (norm : β → ℝ)) {f : α →+* β} (hf : f.IsBounded) (x : α) : norm (f x) ≤ norm x := + contraction_of_isPowMul_of_boundedWrt (SeminormedRing.toRingSeminorm α) hβ hf x + +/-- Given two power-multiplicative ring seminorms `f, g` on `α`, if `f` is bounded by a positive + multiple of `g` and viceversa, then `f = g`. -/ +theorem eq_seminorms {F : Type*} {α : outParam (Type*)} [Ring α] [FunLike F α ℝ] + [RingSeminormClass F α ℝ] {f g : F} (hfpm : IsPowMul f) (hgpm : IsPowMul g) + (hfg : ∃ (r : ℝ) (_ : 0 < r), ∀ a : α, f a ≤ r * g a) + (hgf : ∃ (r : ℝ) (_ : 0 < r), ∀ a : α, g a ≤ r * f a) : f = g := by + obtain ⟨r, hr0, hr⟩ := hfg + obtain ⟨s, hs0, hs⟩ := hgf + have hle : RingHom.IsBoundedWrt f g (RingHom.id _) := ⟨s, hs0, hs⟩ + have hge : RingHom.IsBoundedWrt g f (RingHom.id _) := ⟨r, hr0, hr⟩ + rw [← Function.Injective.eq_iff DFunLike.coe_injective'] + ext x + exact le_antisymm (contraction_of_isPowMul_of_boundedWrt g hfpm hge x) + (contraction_of_isPowMul_of_boundedWrt f hgpm hle x) + +variable {R S : Type*} [NormedCommRing R] [CommRing S] [Algebra R S] + +/-- If `R` is a normed commutative ring and `f₁` and `f₂` are two power-multiplicative `R`-algebra + norms on `S`, then if `f₁` and `f₂` are equivalent on every subring `R[y]` for `y : S`, it + follows that `f₁ = f₂` [BGR, Proposition 3.1.5/1][bosch-guntzer-remmert]. -/ +theorem eq_of_powMul_faithful (f₁ : AlgebraNorm R S) (hf₁_pm : IsPowMul f₁) (f₂ : AlgebraNorm R S) + (hf₂_pm : IsPowMul f₂) + (h_eq : ∀ y : S, ∃ (C₁ C₂ : ℝ) (_ : 0 < C₁) (_ : 0 < C₂), + ∀ x : Algebra.adjoin R {y}, f₁ x.val ≤ C₁ * f₂ x.val ∧ f₂ x.val ≤ C₂ * f₁ x.val) : + f₁ = f₂ := by + ext x + set g₁ : AlgebraNorm R (Algebra.adjoin R ({x} : Set S)) := AlgebraNorm.restriction _ f₁ + set g₂ : AlgebraNorm R (Algebra.adjoin R ({x} : Set S)) := AlgebraNorm.restriction _ f₂ + have hg₁_pm : IsPowMul g₁ := IsPowMul.restriction _ hf₁_pm + have hg₂_pm : IsPowMul g₂ := IsPowMul.restriction _ hf₂_pm + let y : Algebra.adjoin R ({x} : Set S) := ⟨x, Algebra.self_mem_adjoin_singleton R x⟩ + have hy : x = y.val := rfl + have h1 : f₁ y.val = g₁ y := rfl + have h2 : f₂ y.val = g₂ y := rfl + obtain ⟨C₁, C₂, hC₁_pos, hC₂_pos, hC⟩ := h_eq x + obtain ⟨hC₁, hC₂⟩ := forall_and.mp hC + rw [hy, h1, h2, eq_seminorms hg₁_pm hg₂_pm ⟨C₁, hC₁_pos, hC₁⟩ ⟨C₂, hC₂_pos, hC₂⟩] diff --git a/Mathlib/Analysis/Normed/Ring/Seminorm.lean b/Mathlib/Analysis/Normed/Ring/Seminorm.lean index 2f18c14fb0172..650ace74daac1 100644 --- a/Mathlib/Analysis/Normed/Ring/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Ring/Seminorm.lean @@ -151,6 +151,26 @@ end Ring end RingSeminorm +/-- If `f` is a ring seminorm on `a`, then `∀ {n : ℕ}, n ≠ 0 → f (a ^ n) ≤ f a ^ n`. -/ +theorem map_pow_le_pow {F α : Type*} [Ring α] [FunLike F α ℝ] [RingSeminormClass F α ℝ] (f : F) + (a : α) : ∀ {n : ℕ}, n ≠ 0 → f (a ^ n) ≤ f a ^ n + | 0, h => absurd rfl h + | 1, _ => by simp only [pow_one, le_refl] + | n + 2, _ => by + simp only [pow_succ _ (n + 1)]; + exact + le_trans (map_mul_le_mul f _ a) + (mul_le_mul_of_nonneg_right (map_pow_le_pow _ _ n.succ_ne_zero) (apply_nonneg f a)) + +/-- If `f` is a ring seminorm on `a` with `f 1 ≤ 1`, then `∀ (n : ℕ), f (a ^ n) ≤ f a ^ n`. -/ +theorem map_pow_le_pow' {F α : Type*} [Ring α] [FunLike F α ℝ] [RingSeminormClass F α ℝ] {f : F} + (hf1 : f 1 ≤ 1) (a : α) : ∀ n : ℕ, f (a ^ n) ≤ f a ^ n + | 0 => by simp only [pow_zero, hf1] + | n + 1 => by + simp only [pow_succ _ n]; + exact le_trans (map_mul_le_mul f _ a) + (mul_le_mul_of_nonneg_right (map_pow_le_pow' hf1 _ n) (apply_nonneg f a)) + /-- The norm of a `NonUnitalSeminormedRing` as a `RingSeminorm`. -/ def normRingSeminorm (R : Type*) [NonUnitalSeminormedRing R] : RingSeminorm R := { normAddGroupSeminorm R with @@ -368,3 +388,36 @@ lemma MulRingNorm.apply_natAbs_eq {R : Type*} [Ring R] (x : ℤ) (f : MulRingNor f x := by obtain ⟨n, rfl | rfl⟩ := eq_nat_or_neg x <;> simp only [natAbs_neg, natAbs_ofNat, cast_neg, cast_natCast, map_neg_eq_map] + +/-- The seminorm on a `SeminormedRing`, as a `RingSeminorm`. -/ +def SeminormedRing.toRingSeminorm (R : Type*) [SeminormedRing R] : RingSeminorm R where + toFun := norm + map_zero' := norm_zero + add_le' := norm_add_le + mul_le' := norm_mul_le + neg' := norm_neg + +/-- The norm on a `NormedRing`, as a `RingNorm`. -/ +@[simps] +def NormedRing.toRingNorm (R : Type*) [NormedRing R] : RingNorm R where + toFun := norm + map_zero' := norm_zero + add_le' := norm_add_le + mul_le' := norm_mul_le + neg' := norm_neg + eq_zero_of_map_eq_zero' x hx := by rw [← norm_eq_zero]; exact hx + +@[simp] +theorem NormedRing.toRingNorm_apply (R : Type*) [NormedRing R] (x : R) : + (NormedRing.toRingNorm R) x = ‖x‖ := + rfl + +/-- The norm on a `NormedField`, as a `MulRingNorm`. -/ +def NormedField.toMulRingNorm (R : Type*) [NormedField R] : MulRingNorm R where + toFun := norm + map_zero' := norm_zero + map_one' := norm_one + add_le' := norm_add_le + map_mul' := norm_mul + neg' := norm_neg + eq_zero_of_map_eq_zero' x hx := by rw [← norm_eq_zero]; exact hx diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 8de982a5ba82f..7cec019cb7dbe 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -600,3 +600,10 @@ def IsNonarchimedean {A : Type*} [Add A] (f : A → ℝ) : Prop := `f (r ^ n) = (f r) ^ n`. -/ def IsPowMul {R : Type*} [Pow R ℕ] (f : R → ℝ) := ∀ (a : R) {n : ℕ}, 1 ≤ n → f (a ^ n) = f a ^ n + +/-- A ring homomorphism `f : α →+* β` is bounded with respect to the functions `nα : α → ℝ` and + `nβ : β → ℝ` if there exists a positive constant `C` such that for all `x` in `α`, + `nβ (f x) ≤ C * nα x`. -/ +def RingHom.IsBoundedWrt {α : Type*} [Ring α] {β : Type*} [Ring β] (nα : α → ℝ) (nβ : β → ℝ) + (f : α →+* β) : Prop := + ∃ C : ℝ, 0 < C ∧ ∀ x : α, nβ (f x) ≤ C * nα x From c669fd1568a7257d53f7b7e9a17904b555f255bc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 17:17:07 +0000 Subject: [PATCH 125/174] refactor(Data/Real/Archimedean): `sSup`/`sInf` API cleanup (#16735) Add a few lemmas, golf the existing ones and update the docstrings to something more precise --- Mathlib/Analysis/Convex/Gauge.lean | 2 +- Mathlib/Analysis/Normed/Group/Quotient.lean | 2 +- Mathlib/Analysis/Normed/Lp/PiLp.lean | 2 +- .../NormedSpace/Multilinear/Basic.lean | 2 +- .../NormedSpace/OperatorNorm/Basic.lean | 2 +- Mathlib/Analysis/Seminorm.lean | 2 +- Mathlib/Data/NNReal/Basic.lean | 2 +- Mathlib/Data/Real/Archimedean.lean | 234 ++++++++++-------- .../RotationNumber/TranslationNumber.lean | 4 +- 9 files changed, 137 insertions(+), 115 deletions(-) diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index b73b52f91006a..e68cb3085d3d7 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -111,7 +111,7 @@ theorem gauge_of_subset_zero (h : s ⊆ 0) : gauge s = 0 := by /-- The gauge is always nonnegative. -/ theorem gauge_nonneg (x : E) : 0 ≤ gauge s x := - Real.sInf_nonneg _ fun _ hx => hx.1.le + Real.sInf_nonneg fun _ hx => hx.1.le theorem gauge_neg (symmetric : ∀ x ∈ s, -x ∈ s) (x : E) : gauge s (-x) = gauge s x := by have : ∀ x, -x ∈ s ↔ x ∈ s := fun x => ⟨fun h => by simpa using symmetric _ h, symmetric x⟩ diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index 085dab581687e..0759e6a9c2fc6 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -156,7 +156,7 @@ theorem quotient_norm_mk_eq (S : AddSubgroup M) (m : M) : /-- The quotient norm is nonnegative. -/ theorem quotient_norm_nonneg (S : AddSubgroup M) (x : M ⧸ S) : 0 ≤ ‖x‖ := - Real.sInf_nonneg _ <| forall_mem_image.2 fun _ _ ↦ norm_nonneg _ + Real.sInf_nonneg <| forall_mem_image.2 fun _ _ ↦ norm_nonneg _ /-- The quotient norm is nonnegative. -/ theorem norm_mk_nonneg (S : AddSubgroup M) (m : M) : 0 ≤ ‖mk' S m‖ := diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index aed43fb34f51d..cff87b17991f8 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -354,7 +354,7 @@ abbrev pseudoMetricAux : PseudoMetricSpace (PiLp p α) := PseudoMetricSpace.edist_dist] -- Porting note: `le_iSup` needed some help exact le_iSup (fun k => edist (f k) (g k)) i - · refine ENNReal.toReal_le_of_le_ofReal (Real.sSup_nonneg _ ?_) (iSup_le fun i => ?_) + · refine ENNReal.toReal_le_of_le_ofReal (Real.sSup_nonneg ?_) (iSup_le fun i => ?_) · rintro - ⟨i, rfl⟩ exact dist_nonneg · change PseudoMetricSpace.edist _ _ ≤ _ diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 8c50769ed4fbd..c334bfd9e636b 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -339,7 +339,7 @@ theorem isLeast_opNorm : IsLeast {c : ℝ | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * @[deprecated (since := "2024-02-02")] alias isLeast_op_norm := isLeast_opNorm theorem opNorm_nonneg : 0 ≤ ‖f‖ := - Real.sInf_nonneg _ fun _ ⟨hx, _⟩ => hx + Real.sInf_nonneg fun _ ⟨hx, _⟩ => hx @[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index 7c3c4398246cc..a2bd870d5b9d7 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -174,7 +174,7 @@ theorem opNorm_neg (f : E →SL[σ₁₂] F) : ‖-f‖ = ‖f‖ := by simp onl @[deprecated (since := "2024-02-02")] alias op_norm_neg := opNorm_neg theorem opNorm_nonneg (f : E →SL[σ₁₂] F) : 0 ≤ ‖f‖ := - Real.sInf_nonneg _ fun _ ↦ And.left + Real.sInf_nonneg fun _ ↦ And.left @[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index 6ca315beb04f3..c77e95d8e0c87 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -495,7 +495,7 @@ noncomputable instance instSupSet : SupSet (Seminorm 𝕜 E) where if h : BddAbove ((↑) '' s : Set (E → ℝ)) then { toFun := ⨆ p : s, ((p : Seminorm 𝕜 E) : E → ℝ) map_zero' := by - rw [iSup_apply, ← @Real.ciSup_const_zero s] + rw [iSup_apply, ← @Real.iSup_const_zero s] congr! rename_i _ _ _ i exact map_zero i.1 diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index acc5c21138221..da7622c261a6f 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -946,7 +946,7 @@ theorem iInf_empty [IsEmpty ι] (f : ι → ℝ≥0) : ⨅ i, f i = 0 := by @[simp] theorem iInf_const_zero {α : Sort*} : ⨅ _ : α, (0 : ℝ≥0) = 0 := by rw [← coe_inj, coe_iInf] - exact Real.ciInf_const_zero + 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] diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 4df84980d7abb..aab8b86739170 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -17,6 +17,7 @@ open scoped Classical open Pointwise CauSeq namespace Real +variable {ι : Sort*} {f : ι → ℝ} {s : Set ℝ} {a : ℝ} instance instArchimedean : Archimedean ℝ := archimedean_iff_rat_le.2 fun x => @@ -49,9 +50,9 @@ theorem exists_floor (x : ℝ) : ∃ ub : ℤ, (ub : ℝ) ≤ x ∧ ∀ z : ℤ, (let ⟨n, hn⟩ := exists_int_lt x ⟨n, le_of_lt hn⟩) -theorem exists_isLUB {S : Set ℝ} (hne : S.Nonempty) (hbdd : BddAbove S) : ∃ x, IsLUB S x := by +theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x := by rcases hne, hbdd with ⟨⟨L, hL⟩, ⟨U, hU⟩⟩ - have : ∀ d : ℕ, BddAbove { m : ℤ | ∃ y ∈ S, (m : ℝ) ≤ y * d } := by + have : ∀ d : ℕ, BddAbove { m : ℤ | ∃ y ∈ s, (m : ℝ) ≤ y * d } := by cases' exists_int_gt U with k hk refine fun d => ⟨k * d, fun z h => ?_⟩ rcases h with ⟨y, yS, hy⟩ @@ -60,10 +61,10 @@ theorem exists_isLUB {S : Set ℝ} (hne : S.Nonempty) (hbdd : BddAbove S) : ∃ exact mul_le_mul_of_nonneg_right ((hU yS).trans hk.le) d.cast_nonneg choose f hf using fun d : ℕ => Int.exists_greatest_of_bdd (this d) ⟨⌊L * d⌋, L, hL, Int.floor_le _⟩ - have hf₁ : ∀ n > 0, ∃ y ∈ S, ((f n / n : ℚ) : ℝ) ≤ y := fun n n0 => + have hf₁ : ∀ n > 0, ∃ y ∈ s, ((f n / n : ℚ) : ℝ) ≤ y := fun n n0 => let ⟨y, yS, hy⟩ := (hf n).1 ⟨y, yS, by simpa using (div_le_iff₀ (Nat.cast_pos.2 n0 : (_ : ℝ) < _)).2 hy⟩ - have hf₂ : ∀ n > 0, ∀ y ∈ S, (y - ((n : ℕ) : ℝ)⁻¹) < (f n / n : ℚ) := by + have hf₂ : ∀ n > 0, ∀ y ∈ s, (y - ((n : ℕ) : ℝ)⁻¹) < (f n / n : ℚ) := by intro n n0 y yS have := (Int.sub_one_lt_floor _).trans_le (Int.cast_le.2 <| (hf n).2 _ ⟨y, yS, Int.floor_le _⟩) simp only [Rat.cast_div, Rat.cast_intCast, Rat.cast_natCast, gt_iff_lt] @@ -100,56 +101,52 @@ theorem exists_isLUB {S : Set ℝ} (hne : S.Nonempty) (hbdd : BddAbove S) : ∃ le_trans hx (h xS)⟩ /-- A nonempty, bounded below set of real numbers has a greatest lower bound. -/ -theorem exists_isGLB {S : Set ℝ} (hne : S.Nonempty) (hbdd : BddBelow S) : ∃ x, IsGLB S x := by - have hne' : (-S).Nonempty := Set.nonempty_neg.mpr hne - have hbdd' : BddAbove (-S) := bddAbove_neg.mpr hbdd +theorem exists_isGLB (hne : s.Nonempty) (hbdd : BddBelow s) : ∃ x, IsGLB s x := by + have hne' : (-s).Nonempty := Set.nonempty_neg.mpr hne + have hbdd' : BddAbove (-s) := bddAbove_neg.mpr hbdd use -Classical.choose (Real.exists_isLUB hne' hbdd') rw [← isLUB_neg] exact Classical.choose_spec (Real.exists_isLUB hne' hbdd') noncomputable instance : SupSet ℝ := - ⟨fun S => if h : S.Nonempty ∧ BddAbove S then Classical.choose (exists_isLUB h.1 h.2) else 0⟩ + ⟨fun s => if h : s.Nonempty ∧ BddAbove s then Classical.choose (exists_isLUB h.1 h.2) else 0⟩ -theorem sSup_def (S : Set ℝ) : - sSup S = if h : S.Nonempty ∧ BddAbove S then Classical.choose (exists_isLUB h.1 h.2) else 0 := +theorem sSup_def (s : Set ℝ) : + sSup s = if h : s.Nonempty ∧ BddAbove s then Classical.choose (exists_isLUB h.1 h.2) else 0 := rfl -protected theorem isLUB_sSup (S : Set ℝ) (h₁ : S.Nonempty) (h₂ : BddAbove S) : - IsLUB S (sSup S) := by +protected theorem isLUB_sSup (h₁ : s.Nonempty) (h₂ : BddAbove s) : IsLUB s (sSup s) := by simp only [sSup_def, dif_pos (And.intro h₁ h₂)] apply Classical.choose_spec noncomputable instance : InfSet ℝ := - ⟨fun S => -sSup (-S)⟩ + ⟨fun s => -sSup (-s)⟩ -theorem sInf_def (S : Set ℝ) : sInf S = -sSup (-S) := - rfl +theorem sInf_def (s : Set ℝ) : sInf s = -sSup (-s) := rfl -protected theorem is_glb_sInf (S : Set ℝ) (h₁ : S.Nonempty) (h₂ : BddBelow S) : - IsGLB S (sInf S) := by +protected theorem isGLB_sInf (h₁ : s.Nonempty) (h₂ : BddBelow s) : IsGLB s (sInf s) := by rw [sInf_def, ← isLUB_neg', neg_neg] - exact Real.isLUB_sSup _ h₁.neg h₂.neg - -noncomputable instance : ConditionallyCompleteLinearOrder ℝ := - { Real.linearOrder, Real.lattice with - sSup := SupSet.sSup - sInf := InfSet.sInf - le_csSup := fun s a hs ha => (Real.isLUB_sSup s ⟨a, ha⟩ hs).1 ha - csSup_le := fun s a hs ha => (Real.isLUB_sSup s hs ⟨a, ha⟩).2 ha - csInf_le := fun s a hs ha => (Real.is_glb_sInf s ⟨a, ha⟩ hs).1 ha - le_csInf := fun s a hs ha => (Real.is_glb_sInf s hs ⟨a, ha⟩).2 ha - csSup_of_not_bddAbove := fun s hs ↦ by simp [hs, sSup_def] - csInf_of_not_bddBelow := fun s hs ↦ by simp [hs, sInf_def, sSup_def] } - -theorem lt_sInf_add_pos {s : Set ℝ} (h : s.Nonempty) {ε : ℝ} (hε : 0 < ε) : - ∃ a ∈ s, a < sInf s + ε := + exact Real.isLUB_sSup h₁.neg h₂.neg + +@[deprecated (since := "2024-10-02")] alias is_glb_sInf := isGLB_sInf + +noncomputable instance : ConditionallyCompleteLinearOrder ℝ where + __ := Real.linearOrder + __ := Real.lattice + le_csSup s a hs ha := (Real.isLUB_sSup ⟨a, ha⟩ hs).1 ha + csSup_le s a hs ha := (Real.isLUB_sSup hs ⟨a, ha⟩).2 ha + csInf_le s a hs ha := (Real.isGLB_sInf ⟨a, ha⟩ hs).1 ha + le_csInf s a hs ha := (Real.isGLB_sInf hs ⟨a, ha⟩).2 ha + csSup_of_not_bddAbove s hs := by simp [hs, sSup_def] + csInf_of_not_bddBelow s hs := by simp [hs, sInf_def, sSup_def] + +theorem lt_sInf_add_pos (h : s.Nonempty) {ε : ℝ} (hε : 0 < ε) : ∃ a ∈ s, a < sInf s + ε := exists_lt_of_csInf_lt h <| lt_add_of_pos_right _ hε -theorem add_neg_lt_sSup {s : Set ℝ} (h : s.Nonempty) {ε : ℝ} (hε : ε < 0) : - ∃ a ∈ s, sSup s + ε < a := +theorem add_neg_lt_sSup (h : s.Nonempty) {ε : ℝ} (hε : ε < 0) : ∃ a ∈ s, sSup s + ε < a := exists_lt_of_lt_csSup h <| add_lt_iff_neg_left.2 hε -theorem sInf_le_iff {s : Set ℝ} (h : BddBelow s) (h' : s.Nonempty) {a : ℝ} : +theorem sInf_le_iff (h : BddBelow s) (h' : s.Nonempty) : sInf s ≤ a ↔ ∀ ε, 0 < ε → ∃ x ∈ s, x < a + ε := by rw [le_iff_forall_pos_lt_add] constructor <;> intro H ε ε_pos @@ -157,7 +154,7 @@ theorem sInf_le_iff {s : Set ℝ} (h : BddBelow s) (h' : s.Nonempty) {a : ℝ} : · rcases H ε ε_pos with ⟨x, x_in, hx⟩ exact csInf_lt_of_lt h x_in hx -theorem le_sSup_iff {s : Set ℝ} (h : BddAbove s) (h' : s.Nonempty) {a : ℝ} : +theorem le_sSup_iff (h : BddAbove s) (h' : s.Nonempty) : a ≤ sSup s ↔ ∀ ε, ε < 0 → ∃ x ∈ s, a + ε < x := by rw [le_iff_forall_pos_lt_add] refine ⟨fun H ε ε_neg => ?_, fun H ε ε_pos => ?_⟩ @@ -169,102 +166,128 @@ theorem le_sSup_iff {s : Set ℝ} (h : BddAbove s) (h' : s.Nonempty) {a : ℝ} : theorem sSup_empty : sSup (∅ : Set ℝ) = 0 := dif_neg <| by simp -@[simp] lemma iSup_of_isEmpty {α : Sort*} [IsEmpty α] (f : α → ℝ) : ⨆ i, f i = 0 := by +@[simp] lemma iSup_of_isEmpty [IsEmpty ι] (f : ι → ℝ) : ⨆ i, f i = 0 := by dsimp [iSup] convert Real.sSup_empty rw [Set.range_eq_empty_iff] infer_instance @[simp] -theorem ciSup_const_zero {α : Sort*} : ⨆ _ : α, (0 : ℝ) = 0 := by - cases isEmpty_or_nonempty α +theorem iSup_const_zero : ⨆ _ : ι, (0 : ℝ) = 0 := by + cases isEmpty_or_nonempty ι · exact Real.iSup_of_isEmpty _ · exact ciSup_const -theorem sSup_of_not_bddAbove {s : Set ℝ} (hs : ¬BddAbove s) : sSup s = 0 := - dif_neg fun h => hs h.2 - -theorem iSup_of_not_bddAbove {α : Sort*} {f : α → ℝ} (hf : ¬BddAbove (Set.range f)) : - ⨆ i, f i = 0 := - sSup_of_not_bddAbove hf +lemma sSup_of_not_bddAbove (hs : ¬BddAbove s) : sSup s = 0 := dif_neg fun h => hs h.2 +lemma iSup_of_not_bddAbove (hf : ¬BddAbove (Set.range f)) : ⨆ i, f i = 0 := sSup_of_not_bddAbove hf theorem sSup_univ : sSup (@Set.univ ℝ) = 0 := Real.sSup_of_not_bddAbove not_bddAbove_univ @[simp] theorem sInf_empty : sInf (∅ : Set ℝ) = 0 := by simp [sInf_def, sSup_empty] -@[simp] nonrec lemma iInf_of_isEmpty {α : Sort*} [IsEmpty α] (f : α → ℝ) : ⨅ i, f i = 0 := by +@[simp] nonrec lemma iInf_of_isEmpty [IsEmpty ι] (f : ι → ℝ) : ⨅ i, f i = 0 := by rw [iInf_of_isEmpty, sInf_empty] @[simp] -theorem ciInf_const_zero {α : Sort*} : ⨅ _ : α, (0 : ℝ) = 0 := by - cases isEmpty_or_nonempty α +theorem iInf_const_zero : ⨅ _ : ι, (0 : ℝ) = 0 := by + cases isEmpty_or_nonempty ι · exact Real.iInf_of_isEmpty _ · exact ciInf_const -theorem sInf_of_not_bddBelow {s : Set ℝ} (hs : ¬BddBelow s) : sInf s = 0 := +theorem sInf_of_not_bddBelow (hs : ¬BddBelow s) : sInf s = 0 := neg_eq_zero.2 <| sSup_of_not_bddAbove <| mt bddAbove_neg.1 hs -theorem iInf_of_not_bddBelow {α : Sort*} {f : α → ℝ} (hf : ¬BddBelow (Set.range f)) : - ⨅ i, f i = 0 := +theorem iInf_of_not_bddBelow (hf : ¬BddBelow (Set.range f)) : ⨅ i, f i = 0 := sInf_of_not_bddBelow hf -/-- -As `0` is the default value for `Real.sSup` of the empty set or sets which are not bounded above, it -suffices to show that `S` is bounded below by `0` to show that `0 ≤ sSup S`. --/ -theorem sSup_nonneg (S : Set ℝ) (hS : ∀ x ∈ S, (0 : ℝ) ≤ x) : 0 ≤ sSup S := by - rcases S.eq_empty_or_nonempty with (rfl | ⟨y, hy⟩) - · exact sSup_empty.ge - · apply dite _ (fun h => le_csSup_of_le h hy <| hS y hy) fun h => (sSup_of_not_bddAbove h).ge +/-- As `sSup s = 0` when `s` is an empty set of reals, it suffices to show that all elements of `s` +are at most some nonnegative number `a` to show that `sSup s ≤ a`. -/-- -As `0` is the default value for `Real.sSup` of the empty set or sets which are not bounded above, it -suffices to show that `f i` is nonnegative to show that `0 ≤ ⨆ i, f i`. --/ -protected theorem iSup_nonneg {ι : Sort*} {f : ι → ℝ} (hf : ∀ i, 0 ≤ f i) : 0 ≤ ⨆ i, f i := - sSup_nonneg _ <| Set.forall_mem_range.2 hf +See also `csSup_le`. -/ +protected lemma sSup_le (hs : ∀ x ∈ s, x ≤ a) (ha : 0 ≤ a) : sSup s ≤ a := by + obtain rfl | hs' := s.eq_empty_or_nonempty + exacts [sSup_empty.trans_le ha, csSup_le hs' hs] -/-- -As `0` is the default value for `Real.sSup` of the empty set or sets which are not bounded above, it -suffices to show that all elements of `S` are bounded by a nonnegative number to show that `sSup S` -is bounded by this number. --/ -protected theorem sSup_le {S : Set ℝ} {a : ℝ} (hS : ∀ x ∈ S, x ≤ a) (ha : 0 ≤ a) : sSup S ≤ a := by - rcases S.eq_empty_or_nonempty with (rfl | hS₂) - exacts [sSup_empty.trans_le ha, csSup_le hS₂ hS] +/-- As `⨆ i, f i = 0` when the domain of the real-valued function `f` is empty, it suffices to show +that all values of `f` are at most some nonnegative number `a` to show that `⨆ i, f i ≤ a`. -protected theorem iSup_le {ι : Sort*} {f : ι → ℝ} {a : ℝ} (hS : ∀ i, f i ≤ a) (ha : 0 ≤ a) : - ⨆ i, f i ≤ a := - Real.sSup_le (Set.forall_mem_range.2 hS) ha +See also `ciSup_le`. -/ +protected lemma iSup_le (hf : ∀ i, f i ≤ a) (ha : 0 ≤ a) : ⨆ i, f i ≤ a := + Real.sSup_le (Set.forall_mem_range.2 hf) ha -/-- As `0` is the default value for `Real.sSup` of the empty set, it suffices to show that `S` is -bounded above by `0` to show that `sSup S ≤ 0`. --/ -theorem sSup_nonpos (S : Set ℝ) (hS : ∀ x ∈ S, x ≤ (0 : ℝ)) : sSup S ≤ 0 := - Real.sSup_le hS le_rfl +/-- As `sInf s = 0` when `s` is an empty set of reals, it suffices to show that all elements of `s` +are at least some nonpositive number `a` to show that `a ≤ sInf s`. -/-- As `0` is the default value for `Real.sInf` of the empty set, it suffices to show that `S` is -bounded below by `0` to show that `0 ≤ sInf S`. --/ -theorem sInf_nonneg (S : Set ℝ) (hS : ∀ x ∈ S, (0 : ℝ) ≤ x) : 0 ≤ sInf S := by - rcases S.eq_empty_or_nonempty with (rfl | hS₂) - exacts [sInf_empty.ge, le_csInf hS₂ hS] +See also `le_csInf`. -/ +protected lemma le_sInf (hs : ∀ x ∈ s, a ≤ x) (ha : a ≤ 0) : a ≤ sInf s := by + obtain rfl | hs' := s.eq_empty_or_nonempty + exacts [ha.trans_eq sInf_empty.symm, le_csInf hs' hs] -/-- As `0` is the default value for `Real.sInf` of the empty set, it suffices to show that `f i` is -bounded below by `0` to show that `0 ≤ iInf f`. --/ -theorem iInf_nonneg {ι} {f : ι → ℝ} (hf : ∀ i, 0 ≤ f i) : 0 ≤ iInf f := - sInf_nonneg _ <| Set.forall_mem_range.2 hf +/-- As `⨅ i, f i = 0` when the domain of the real-valued function `f` is empty, it suffices to show +that all values of `f` are at least some nonpositive number `a` to show that `a ≤ ⨅ i, f i`. -/-- -As `0` is the default value for `Real.sInf` of the empty set or sets which are not bounded below, it -suffices to show that `S` is bounded above by `0` to show that `sInf S ≤ 0`. --/ -theorem sInf_nonpos (S : Set ℝ) (hS : ∀ x ∈ S, x ≤ (0 : ℝ)) : sInf S ≤ 0 := by - rcases S.eq_empty_or_nonempty with (rfl | ⟨y, hy⟩) +See also `le_ciInf`. -/ +protected lemma le_iInf (hf : ∀ i, a ≤ f i) (ha : a ≤ 0) : a ≤ ⨅ i, f i := + Real.le_sInf (Set.forall_mem_range.2 hf) ha + +/-- As `sSup s = 0` when `s` is an empty set of reals, it suffices to show that all elements of `s` +are nonpositive to show that `sSup s ≤ 0`. -/ +lemma sSup_nonpos (hs : ∀ x ∈ s, x ≤ 0) : sSup s ≤ 0 := Real.sSup_le hs le_rfl + +/-- As `⨆ i, f i = 0` when the domain of the real-valued function `f` is empty, +it suffices to show that all values of `f` are nonpositive to show that `⨆ i, f i ≤ 0`. -/ +lemma iSup_nonpos (hf : ∀ i, f i ≤ 0) : ⨆ i, f i ≤ 0 := Real.iSup_le hf le_rfl + +/-- As `sInf s = 0` when `s` is an empty set of reals, it suffices to show that all elements of `s` +are nonnegative to show that `0 ≤ sInf s`. -/ +lemma sInf_nonneg (hs : ∀ x ∈ s, 0 ≤ x) : 0 ≤ sInf s := Real.le_sInf hs le_rfl + +/-- As `⨅ i, f i = 0` when the domain of the real-valued function `f` is empty, +it suffices to show that all values of `f` are nonnegative to show that `0 ≤ ⨅ i, f i`. -/ +lemma iInf_nonneg (hf : ∀ i, 0 ≤ f i) : 0 ≤ iInf f := Real.le_iInf hf le_rfl + +/-- As `sSup s = 0` when `s` is a set of reals that's unbounded above, it suffices to show that `s` +contains a nonnegative element to show that `0 ≤ sSup s`. -/ +lemma sSup_nonneg' (hs : ∃ x ∈ s, 0 ≤ x) : 0 ≤ sSup s := by + obtain ⟨x, hxs, hx⟩ := hs + exact dite _ (fun h ↦ le_csSup_of_le h hxs hx) fun h ↦ (sSup_of_not_bddAbove h).ge + +/-- As `⨆ i, f i = 0` when the real-valued function `f` is unbounded above, +it suffices to show that `f` takes a nonnegative value to show that `0 ≤ ⨆ i, f i`. -/ +lemma iSup_nonneg' (hf : ∃ i, 0 ≤ f i) : 0 ≤ ⨆ i, f i := sSup_nonneg' <| Set.exists_range_iff.2 hf + +/-- As `sInf s = 0` when `s` is a set of reals that's unbounded below, it suffices to show that `s` +contains a nonpositive element to show that `sInf s ≤ 0`. -/ +lemma sInf_nonpos' (hs : ∃ x ∈ s, x ≤ 0) : sInf s ≤ 0 := by + obtain ⟨x, hxs, hx⟩ := hs + exact dite _ (fun h ↦ csInf_le_of_le h hxs hx) fun h ↦ (sInf_of_not_bddBelow h).le + +/-- As `⨅ i, f i = 0` when the real-valued function `f` is unbounded below, +it suffices to show that `f` takes a nonpositive value to show that `0 ≤ ⨅ i, f i`. -/ +lemma iInf_nonpos' (hf : ∃ i, f i ≤ 0) : ⨅ i, f i ≤ 0 := sInf_nonpos' <| Set.exists_range_iff.2 hf + +/-- As `sSup s = 0` when `s` is a set of reals that's either empty or unbounded above, +it suffices to show that all elements of `s` are nonnegative to show that `0 ≤ sSup s`. -/ +lemma sSup_nonneg (hs : ∀ x ∈ s, 0 ≤ x) : 0 ≤ sSup s := by + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty + · exact sSup_empty.ge + · exact sSup_nonneg' ⟨x, hx, hs _ hx⟩ + +/-- As `⨆ i, f i = 0` when the domain of the real-valued function `f` is empty or unbounded above, +it suffices to show that all values of `f` are nonnegative to show that `0 ≤ ⨆ i, f i`. -/ +lemma iSup_nonneg (hf : ∀ i, 0 ≤ f i) : 0 ≤ ⨆ i, f i := sSup_nonneg <| Set.forall_mem_range.2 hf + +/-- As `sInf s = 0` when `s` is a set of reals that's either empty or unbounded below, +it suffices to show that all elements of `s` are nonpositive to show that `sInf s ≤ 0`. -/ +lemma sInf_nonpos (hs : ∀ x ∈ s, x ≤ 0) : sInf s ≤ 0 := by + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty · exact sInf_empty.le - · apply dite _ (fun h => csInf_le_of_le h hy <| hS y hy) fun h => (sInf_of_not_bddBelow h).le + · exact sInf_nonpos' ⟨x, hx, hs _ hx⟩ + +/-- As `⨅ i, f i = 0` when the domain of the real-valued function `f` is empty or unbounded below, +it suffices to show that all values of `f` are nonpositive to show that `0 ≤ ⨅ i, f i`. -/ +lemma iInf_nonpos (hf : ∀ i, f i ≤ 0) : ⨅ i, f i ≤ 0 := sInf_nonpos <| Set.forall_mem_range.2 hf theorem sInf_le_sSup (s : Set ℝ) (h₁ : BddBelow s) (h₂ : BddAbove s) : sInf s ≤ sSup s := by rcases s.eq_empty_or_nonempty with (rfl | hne) @@ -272,12 +295,12 @@ theorem sInf_le_sSup (s : Set ℝ) (h₁ : BddBelow s) (h₂ : BddAbove s) : sIn · exact csInf_le_csSup h₁ h₂ hne theorem cauSeq_converges (f : CauSeq ℝ abs) : ∃ x, f ≈ const abs x := by - let S := { x : ℝ | const abs x < f } - have lb : ∃ x, x ∈ S := exists_lt f - have ub' : ∀ x, f < const abs x → ∀ y ∈ S, y ≤ x := fun x h y yS => + let s := {x : ℝ | const abs x < f} + have lb : ∃ x, x ∈ s := exists_lt f + have ub' : ∀ x, f < const abs x → ∀ y ∈ s, y ≤ x := fun x h y yS => le_of_lt <| const_lt.1 <| CauSeq.lt_trans yS h - have ub : ∃ x, ∀ y ∈ S, y ≤ x := (exists_gt f).imp ub' - refine ⟨sSup S, ((lt_total _ _).resolve_left fun h => ?_).resolve_right fun h => ?_⟩ + have ub : ∃ x, ∀ y ∈ s, y ≤ x := (exists_gt f).imp ub' + refine ⟨sSup s, ((lt_total _ _).resolve_left fun h => ?_).resolve_right fun h => ?_⟩ · rcases h with ⟨ε, ε0, i, ih⟩ refine (csSup_le lb (ub' _ ?_)).not_lt (sub_lt_self _ (half_pos ε0)) refine ⟨_, half_pos ε0, i, fun j ij => ?_⟩ @@ -336,8 +359,7 @@ theorem iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by exact iInter_Iic_eq_empty_iff.mpr not_bddBelow_coe /-- Exponentiation is eventually larger than linear growth. -/ -lemma exists_natCast_add_one_lt_pow_of_one_lt {a : ℝ} (ha : 1 < a) : - ∃ m : ℕ, (m + 1 : ℝ) < a ^ m := by +lemma exists_natCast_add_one_lt_pow_of_one_lt (ha : 1 < a) : ∃ m : ℕ, (m + 1 : ℝ) < a ^ m := by obtain ⟨k, posk, hk⟩ : ∃ k : ℕ, 0 < k ∧ 1 / k + 1 < a := by contrapose! ha refine le_of_forall_lt_rat_imp_le ?_ diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index a47e9d3145ffb..bf93db2355a6d 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -852,8 +852,8 @@ theorem semiconj_of_group_action_of_forall_translationNumber_eq {G : Type*} [Gro have hF₁ : ∀ g, ⇑(F₁ g) = f₁ g := fun _ => rfl have hF₂ : ∀ g, ⇑(F₂ g) = f₂ g := fun _ => rfl -- Now we apply `csSup_div_semiconj` and go back to `f₁` and `f₂`. - refine ⟨⟨⟨_, fun x y hxy => ?_⟩, fun x => ?_⟩, csSup_div_semiconj F₂ F₁ fun x => ?_⟩ <;> - simp only [hF₁, hF₂, ← map_inv, coe_mk] + refine ⟨⟨⟨fun x ↦ ⨆ g', (F₂ g')⁻¹ (F₁ g' x), fun x y hxy => ?_⟩, fun x => ?_⟩, + csSup_div_semiconj F₂ F₁ fun x => ?_⟩ <;> simp only [hF₁, hF₂, ← map_inv, coe_mk] · exact ciSup_mono (this y) fun g => mono _ (mono _ hxy) · simp only [map_add_one] exact (Monotone.map_ciSup_of_continuousAt (continuousAt_id.add continuousAt_const) From b490daf06ae807dea5b103912bdded77e46ccf4a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 17:17:09 +0000 Subject: [PATCH 126/174] chore: make `Submodule.span_mono` be `gcongr` (#17293) From LeanCamCombi --- Mathlib/LinearAlgebra/Span.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 8ff62644c2b01..460f7003ad94b 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -74,7 +74,7 @@ theorem subset_span : s ⊆ span R s := fun _ h => mem_span.2 fun _ hp => hp h theorem span_le {p} : span R s ≤ p ↔ s ⊆ p := ⟨Subset.trans subset_span, fun ss _ h => mem_span.1 h _ ss⟩ -theorem span_mono (h : s ⊆ t) : span R s ≤ span R t := +@[gcongr] theorem span_mono (h : s ⊆ t) : span R s ≤ span R t := span_le.2 <| Subset.trans h subset_span theorem span_monotone : Monotone (span R : Set M → Submodule R M) := fun _ _ => span_mono From 2aa61a8c55cb42718c0221f817c3f515e6c0833d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 17:17:10 +0000 Subject: [PATCH 127/174] feat: `s / t` is finite (#17306) From PFR --- Mathlib/Data/Set/Pointwise/Finite.lean | 63 +++++++++++++++++--------- 1 file changed, 41 insertions(+), 22 deletions(-) diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Data/Set/Pointwise/Finite.lean index 934fd477ecac1..370ad1550f6b6 100644 --- a/Mathlib/Data/Set/Pointwise/Finite.lean +++ b/Mathlib/Data/Set/Pointwise/Finite.lean @@ -26,16 +26,6 @@ theorem finite_one : (1 : Set α).Finite := end One -section InvolutiveInv - -variable [InvolutiveInv α] {s : Set α} - -@[to_additive] -theorem Finite.inv (hs : s.Finite) : s⁻¹.Finite := - hs.preimage inv_injective.injOn - -end InvolutiveInv - section Mul variable [Mul α] {s t : Set α} @@ -109,20 +99,52 @@ section Cancel variable [Mul α] [IsLeftCancelMul α] [IsRightCancelMul α] {s t : Set α} @[to_additive] -theorem infinite_mul : (s * t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := - infinite_image2 (fun _ _ => (mul_left_injective _).injOn) fun _ _ => - (mul_right_injective _).injOn +lemma finite_mul : (s * t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := + finite_image2 (fun _ _ ↦ (mul_left_injective _).injOn) fun _ _ ↦ (mul_right_injective _).injOn @[to_additive] -lemma finite_mul : (s * t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := - finite_image2 (fun _ _ ↦ (mul_left_injective _).injOn) - fun _ _ ↦ (mul_right_injective _).injOn +lemma infinite_mul : (s * t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := + infinite_image2 (fun _ _ => (mul_left_injective _).injOn) fun _ _ => (mul_right_injective _).injOn end Cancel +section InvolutiveInv +variable [InvolutiveInv α] {s : Set α} + +@[to_additive (attr := simp)] lemma finite_inv : s⁻¹.Finite ↔ s.Finite := by + rw [← image_inv, finite_image_iff inv_injective.injOn] + +@[to_additive (attr := simp)] lemma infinite_inv : s⁻¹.Infinite ↔ s.Infinite := finite_inv.not + +@[to_additive] alias ⟨Finite.of_inv, Finite.inv⟩ := finite_inv + +end InvolutiveInv + +section Div +variable [Div α] {s t : Set α} + +@[to_additive] lemma Finite.div : s.Finite → t.Finite → (s / t).Finite := .image2 _ + +/-- Division preserves finiteness. -/ +@[to_additive "Subtraction preserves finiteness."] +def fintypeDiv [DecidableEq α] (s t : Set α) [Fintype s] [Fintype t] : Fintype (s / t) := + Set.fintypeImage2 _ _ _ + +end Div + section Group -variable [Group α] [MulAction α β] {a : α} {s : Set β} +variable [Group α] {s t : Set α} + +@[to_additive] +lemma finite_div : (s / t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := + finite_image2 (fun _ _ ↦ div_left_injective.injOn) fun _ _ ↦ div_right_injective.injOn + +@[to_additive] +lemma infinite_div : (s / t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := + infinite_image2 (fun _ _ ↦ div_left_injective.injOn) fun _ _ ↦ div_right_injective.injOn + +variable [MulAction α β] {a : α} {s : Set β} @[to_additive (attr := simp)] theorem finite_smul_set : (a • s).Finite ↔ s.Finite := @@ -132,11 +154,8 @@ theorem finite_smul_set : (a • s).Finite ↔ s.Finite := theorem infinite_smul_set : (a • s).Infinite ↔ s.Infinite := infinite_image_iff (MulAction.injective _).injOn -alias ⟨Finite.of_smul_set, _⟩ := finite_smul_set - -alias ⟨_, Infinite.smul_set⟩ := infinite_smul_set - -attribute [to_additive] Finite.of_smul_set Infinite.smul_set +@[to_additive] alias ⟨Finite.of_smul_set, _⟩ := finite_smul_set +@[to_additive] alias ⟨_, Infinite.smul_set⟩ := infinite_smul_set end Group From 27899cdbc763c34501196f9effe26ca33fca78ff Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Wed, 2 Oct 2024 17:17:11 +0000 Subject: [PATCH 128/174] feat(Algebra/Polynomial/FieldDivision): `natDegree_mod_lt` (#17318) --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 8f5cad434faf7..06134da2f5675 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -382,6 +382,21 @@ theorem map_mod [Field k] (f : R →+* k) : (p % q).map f = p.map f % q.map f := · rw [mod_def, mod_def, leadingCoeff_map f, ← map_inv₀ f, ← map_C f, ← Polynomial.map_mul f, map_modByMonic f (monic_mul_leadingCoeff_inv hq0)] +lemma natDegree_mod_lt [Field k] (p : k[X]) {q : k[X]} (hq : q.natDegree ≠ 0) : + (p % q).natDegree < q.natDegree := by + have hq' : q.leadingCoeff ≠ 0 := by + rw [leadingCoeff_ne_zero] + contrapose! hq + simp [hq] + rw [mod_def] + refine (natDegree_modByMonic_lt p ?_ ?_).trans_le ?_ + · refine monic_mul_C_of_leadingCoeff_mul_eq_one ?_ + rw [mul_inv_eq_one₀ hq'] + · contrapose! hq + rw [← natDegree_mul_C_eq_of_mul_eq_one ((inv_mul_eq_one₀ hq').mpr rfl)] + simp [hq] + · exact natDegree_mul_C_le q q.leadingCoeff⁻¹ + section open EuclideanDomain From 24484cb75991650b09089656bf79a86fc17c22cc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 17:17:12 +0000 Subject: [PATCH 129/174] feat(Subgraph): API lemmas (#17332) From LeanCamCombi --- .../Combinatorics/SimpleGraph/Subgraph.lean | 71 +++++++++++++------ 1 file changed, 50 insertions(+), 21 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 9d64a2270e2d0..c0cc9e7d6ca21 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -133,6 +133,9 @@ theorem coe_adj_sub (G' : Subgraph G) (u v : G'.verts) (h : G'.coe.Adj u v) : G. protected theorem Adj.coe {H : G.Subgraph} {u v : V} (h : H.Adj u v) : H.coe.Adj ⟨u, H.edge_vert h⟩ ⟨v, H.edge_vert h.symm⟩ := h +instance (G : SimpleGraph V) (H : Subgraph G) [DecidableRel H.Adj] : DecidableRel H.coe.Adj := + fun a b ↦ ‹DecidableRel H.Adj› _ _ + /-- A subgraph is called a *spanning subgraph* if it contains all the vertices of `G`. -/ def IsSpanning (G' : Subgraph G) : Prop := ∀ v : V, v ∈ G'.verts @@ -156,6 +159,8 @@ theorem Adj.of_spanningCoe {G' : Subgraph G} {u v : G'.verts} (h : G'.spanningCo G.Adj u v := G'.adj_sub h +lemma spanningCoe_le (G' : G.Subgraph) : G'.spanningCoe ≤ G := fun _ _ ↦ G'.3 + theorem spanningCoe_inj : G₁.spanningCoe = G₂.spanningCoe ↔ G₁.Adj = G₂.Adj := by simp [Subgraph.spanningCoe] @@ -209,15 +214,27 @@ theorem edgeSet_subset (G' : Subgraph G) : G'.edgeSet ⊆ G.edgeSet := Sym2.ind (fun _ _ ↦ G'.adj_sub) @[simp] -theorem mem_edgeSet {G' : Subgraph G} {v w : V} : s(v, w) ∈ G'.edgeSet ↔ G'.Adj v w := Iff.rfl +protected lemma mem_edgeSet {G' : Subgraph G} {v w : V} : s(v, w) ∈ G'.edgeSet ↔ G'.Adj v w := .rfl + +@[simp] lemma edgeSet_coe {G' : G.Subgraph} : G'.coe.edgeSet = Sym2.map (↑) ⁻¹' G'.edgeSet := by + ext e; induction' e using Sym2.ind with a b; simp -theorem mem_verts_if_mem_edge {G' : Subgraph G} {e : Sym2 V} {v : V} (he : e ∈ G'.edgeSet) +lemma image_coe_edgeSet_coe (G' : G.Subgraph) : Sym2.map (↑) '' G'.coe.edgeSet = G'.edgeSet := by + rw [edgeSet_coe, Set.image_preimage_eq_iff] + rintro e he + induction' e using Sym2.ind with a b + rw [Subgraph.mem_edgeSet] at he + exact ⟨s(⟨a, edge_vert _ he⟩, ⟨b, edge_vert _ he.symm⟩), Sym2.map_pair_eq ..⟩ + +theorem mem_verts_of_mem_edge {G' : Subgraph G} {e : Sym2 V} {v : V} (he : e ∈ G'.edgeSet) (hv : v ∈ e) : v ∈ G'.verts := by induction e rcases Sym2.mem_iff.mp hv with (rfl | rfl) · exact G'.edge_vert he · exact G'.edge_vert (G'.symm he) +@[deprecated (since := "2024-10-01")] alias mem_verts_if_mem_edge := mem_verts_of_mem_edge + /-- The `incidenceSet` is the set of edges incident to a given vertex. -/ def incidenceSet (G' : Subgraph G) (v : V) : Set (Sym2 V) := {e ∈ G'.edgeSet | v ∈ e} @@ -377,6 +394,18 @@ theorem verts_iSup {f : ι → G.Subgraph} : (⨆ i, f i).verts = ⋃ i, (f i).v @[simp] theorem verts_iInf {f : ι → G.Subgraph} : (⨅ i, f i).verts = ⋂ i, (f i).verts := by simp [iInf] +@[simp] lemma coe_bot : (⊥ : G.Subgraph).coe = ⊥ := rfl + +@[simp] lemma IsInduced.top : (⊤ : G.Subgraph).IsInduced := fun _ _ ↦ id + +/-- The graph isomorphism between the top element of `G.subgraph` and `G`. -/ +def topIso : (⊤ : G.Subgraph).coe ≃g G where + toFun := (↑) + invFun a := ⟨a, Set.mem_univ _⟩ + left_inv _ := Subtype.eta .. + right_inv _ := rfl + map_rel_iff' := .rfl + theorem verts_spanningCoe_injective : (fun G' : Subgraph G => (G'.verts, G'.spanningCoe)).Injective := by intro G₁ G₂ h @@ -551,9 +580,12 @@ theorem _root_.Disjoint.edgeSet {H₁ H₂ : Subgraph G} (h : Disjoint H₁ H₂ Disjoint H₁.edgeSet H₂.edgeSet := disjoint_iff_inf_le.mpr <| by simpa using edgeSet_mono h.le_bot +section map +variable {G' : SimpleGraph W} {f : G →g G'} + /-- Graph homomorphisms induce a covariant function on subgraphs. -/ @[simps] -protected def map {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) : G'.Subgraph where +protected def map (f : G →g G') (H : G.Subgraph) : G'.Subgraph where verts := f '' H.verts Adj := Relation.Map H.Adj f f adj_sub := by @@ -566,29 +598,26 @@ protected def map {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph) : G'.Sub rintro _ _ ⟨u, v, h, rfl, rfl⟩ exact ⟨v, u, H.symm h, rfl, rfl⟩ -theorem map_monotone {G' : SimpleGraph W} (f : G →g G') : Monotone (Subgraph.map f) := by - intro H H' h +@[simp] lemma map_id (H : G.Subgraph) : H.map Hom.id = H := by ext <;> simp + +lemma map_comp {U : Type*} {G'' : SimpleGraph U} (H : G.Subgraph) (f : G →g G') (g : G' →g G'') : + H.map (g.comp f) = (H.map f).map g := by ext <;> simp [Subgraph.map] + +@[gcongr] lemma map_mono {H₁ H₂ : G.Subgraph} (hH : H₁ ≤ H₂) : H₁.map f ≤ H₂.map f := by constructor · intro simp only [map_verts, Set.mem_image, forall_exists_index, and_imp] rintro v hv rfl - exact ⟨_, h.1 hv, rfl⟩ + exact ⟨_, hH.1 hv, rfl⟩ · rintro _ _ ⟨u, v, ha, rfl, rfl⟩ - exact ⟨_, _, h.2 ha, rfl, rfl⟩ - -theorem map_sup {G : SimpleGraph V} {G' : SimpleGraph W} (f : G →g G') {H H' : G.Subgraph} : - (H ⊔ H').map f = H.map f ⊔ H'.map f := by - ext1 - · simp only [Set.image_union, map_verts, verts_sup] - · ext - simp only [Relation.Map, map_adj, sup_adj] - constructor - · rintro ⟨a, b, h | h, rfl, rfl⟩ - · exact Or.inl ⟨_, _, h, rfl, rfl⟩ - · exact Or.inr ⟨_, _, h, rfl, rfl⟩ - · rintro (⟨a, b, h, rfl, rfl⟩ | ⟨a, b, h, rfl, rfl⟩) - · exact ⟨_, _, Or.inl h, rfl, rfl⟩ - · exact ⟨_, _, Or.inr h, rfl, rfl⟩ + exact ⟨_, _, hH.2 ha, rfl, rfl⟩ + +lemma map_monotone : Monotone (Subgraph.map f) := fun _ _ ↦ map_mono + +theorem map_sup (f : G →g G') (H₁ H₂ : G.Subgraph) : (H₁ ⊔ H₂).map f = H₁.map f ⊔ H₂.map f := by + ext <;> simp [Set.image_union, map_adj, sup_adj, Relation.Map, or_and_right, exists_or] + +end map /-- Graph homomorphisms induce a contravariant function on subgraphs. -/ @[simps] From 5cb57a1ae7d060bb15e87afbe8d01d0a7bbf52d1 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 2 Oct 2024 17:17:13 +0000 Subject: [PATCH 130/174] =?UTF-8?q?feat:=20golf=20proof=20of=20Fr=C3=A9che?= =?UTF-8?q?t=E2=80=93von=20Neumann=E2=80=93Jordan=20theorem=20(#17349)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Exploit algebraic automation, including the new `module` tactic, more heavily in the proof of the Fréchet–von Neumann–Jordan theorem (i.e., the theorem that a norm satisfying the parallellogram identity comes from an inner product). --- .../Analysis/InnerProductSpace/OfNorm.lean | 201 ++++++------------ 1 file changed, 61 insertions(+), 140 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index 48732656917c7..b29417e75a1af 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -5,6 +5,8 @@ Authors: Heather Macbeth -/ import Mathlib.Topology.Algebra.Algebra import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Algebra.Module.LinearMap.Rat +import Mathlib.Tactic.Module /-! # Inner product space derived from a norm @@ -97,173 +99,97 @@ private def innerProp' (r : 𝕜) : Prop := variable {E} -theorem innerProp_neg_one : innerProp' E ((-1 : ℤ) : 𝕜) := by - intro x y - simp only [inner_, neg_mul_eq_neg_mul, one_mul, Int.cast_one, one_smul, RingHom.map_one, map_neg, - Int.cast_neg, neg_smul, neg_one_mul] - rw [neg_mul_comm] - congr 1 - have h₁ : ‖-x - y‖ = ‖x + y‖ := by rw [← neg_add', norm_neg] - have h₂ : ‖-x + y‖ = ‖x - y‖ := by rw [← neg_sub, norm_neg, sub_eq_neg_add] - have h₃ : ‖(I : 𝕜) • -x + y‖ = ‖(I : 𝕜) • x - y‖ := by - rw [← neg_sub, norm_neg, sub_eq_neg_add, ← smul_neg] - have h₄ : ‖(I : 𝕜) • -x - y‖ = ‖(I : 𝕜) • x + y‖ := by rw [smul_neg, ← neg_add', norm_neg] - rw [h₁, h₂, h₃, h₄] - ring - theorem _root_.Continuous.inner_ {f g : ℝ → E} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => inner_ 𝕜 (f x) (g x) := by unfold inner_ fun_prop theorem inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := by - simp only [inner_] - have h₁ : RCLike.normSq (4 : 𝕜) = 16 := by - have : ((4 : ℝ) : 𝕜) = (4 : 𝕜) := by norm_cast - rw [← this, normSq_eq_def', RCLike.norm_of_nonneg (by norm_num : (0 : ℝ) ≤ 4)] - norm_num - have h₂ : ‖x + x‖ = 2 * ‖x‖ := by rw [← two_smul 𝕜, norm_smul, RCLike.norm_two] - simp only [h₁, h₂, algebraMap_eq_ofReal, sub_self, norm_zero, mul_re, inv_re, ofNat_re, map_sub, - map_add, ofReal_re, ofNat_im, ofReal_im, mul_im, I_re, inv_im] + simp only [inner_, normSq_apply, ofNat_re, ofNat_im, map_sub, map_add, map_zero, map_mul, + ofReal_re, ofReal_im, mul_re, inv_re, mul_im, I_re, inv_im] + have h₁ : ‖x - x‖ = 0 := by simp + have h₂ : ‖x + x‖ = 2 • ‖x‖ := by convert norm_nsmul 𝕜 2 x using 2; module + rw [h₁, h₂] ring -attribute [local simp] map_ofNat in -- use `ofNat` simp theorem with bad keys theorem inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := by - simp only [inner_] - have h4 : conj (4⁻¹ : 𝕜) = 4⁻¹ := by norm_num - rw [map_mul, h4] - congr 1 - simp only [map_sub, map_add, conj_ofReal, map_mul, conj_I] + simp only [inner_, map_sub, map_add, map_mul, map_inv₀, map_ofNat, conj_ofReal, conj_I] rw [add_comm y x, norm_sub_rev] by_cases hI : (I : 𝕜) = 0 · simp only [hI, neg_zero, zero_mul] - -- Porting note: this replaces `norm_I_of_ne_zero` which does not exist in Lean 4 - have : ‖(I : 𝕜)‖ = 1 := by - rw [← mul_self_inj_of_nonneg (norm_nonneg I) zero_le_one, one_mul, ← norm_mul, - I_mul_I_of_nonzero hI, norm_neg, norm_one] + have hI' := I_mul_I_of_nonzero hI + have I_smul (v : E) : ‖(I : 𝕜) • v‖ = ‖v‖ := by rw [norm_smul, norm_I_of_ne_zero hI, one_mul] have h₁ : ‖(I : 𝕜) • y - x‖ = ‖(I : 𝕜) • x + y‖ := by - trans ‖(I : 𝕜) • ((I : 𝕜) • y - x)‖ - · rw [norm_smul, this, one_mul] - · rw [smul_sub, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ← neg_add', add_comm, norm_neg] + convert I_smul ((I : 𝕜) • x + y) using 2 + linear_combination (norm := module) congr(-$hI' • x) have h₂ : ‖(I : 𝕜) • y + x‖ = ‖(I : 𝕜) • x - y‖ := by - trans ‖(I : 𝕜) • ((I : 𝕜) • y + x)‖ - · rw [norm_smul, this, one_mul] - · rw [smul_add, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ← neg_add_eq_sub] - rw [h₁, h₂, ← sub_add_eq_add_sub] - simp only [neg_mul, sub_eq_add_neg, neg_neg] + convert (I_smul ((I : 𝕜) • y + x)).symm using 2 + linear_combination (norm := module) congr(-$hI' • y) + rw [h₁, h₂] + ring variable [InnerProductSpaceable E] -private theorem add_left_aux1 (x y z : E) : ‖x + y + z‖ * ‖x + y + z‖ = - (‖2 • x + y‖ * ‖2 • x + y‖ + ‖2 • z + y‖ * ‖2 • z + y‖) / 2 - ‖x - z‖ * ‖x - z‖ := by - rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm] - convert parallelogram_identity (x + y + z) (x - z) using 4 <;> · rw [two_smul]; abel - -private theorem add_left_aux2 (x y z : E) : ‖x + y - z‖ * ‖x + y - z‖ = - (‖2 • x + y‖ * ‖2 • x + y‖ + ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 - ‖x + z‖ * ‖x + z‖ := by - rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm] - have h₀ := parallelogram_identity (x + y - z) (x + z) - convert h₀ using 4 <;> · rw [two_smul]; abel +private theorem add_left_aux1 (x y z : E) : + ‖2 • x + y‖ * ‖2 • x + y‖ + ‖2 • z + y‖ * ‖2 • z + y‖ + = 2 * (‖x + y + z‖ * ‖x + y + z‖ + ‖x - z‖ * ‖x - z‖) := by + convert parallelogram_identity (x + y + z) (x - z) using 4 <;> abel -private theorem add_left_aux2' (x y z : E) : - ‖x + y + z‖ * ‖x + y + z‖ - ‖x + y - z‖ * ‖x + y - z‖ = - ‖x + z‖ * ‖x + z‖ - ‖x - z‖ * ‖x - z‖ + - (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 := by - rw [add_left_aux1, add_left_aux2]; ring +private theorem add_left_aux2 (x y z : E) : ‖2 • x + y‖ * ‖2 • x + y‖ + ‖y - 2 • z‖ * ‖y - 2 • z‖ + = 2 * (‖x + y - z‖ * ‖x + y - z‖ + ‖x + z‖ * ‖x + z‖) := by + convert parallelogram_identity (x + y - z) (x + z) using 4 <;> abel private theorem add_left_aux3 (y z : E) : - ‖2 • z + y‖ * ‖2 • z + y‖ = 2 * (‖y + z‖ * ‖y + z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := by - apply eq_sub_of_add_eq - convert parallelogram_identity (y + z) z using 4 <;> (try rw [two_smul]) <;> abel + ‖2 • z + y‖ * ‖2 • z + y‖ + ‖y‖ * ‖y‖ = 2 * (‖y + z‖ * ‖y + z‖ + ‖z‖ * ‖z‖) := by + convert parallelogram_identity (y + z) z using 4 <;> abel private theorem add_left_aux4 (y z : E) : - ‖y - 2 • z‖ * ‖y - 2 • z‖ = 2 * (‖y - z‖ * ‖y - z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := by - apply eq_sub_of_add_eq' - have h₀ := parallelogram_identity (y - z) z - convert h₀ using 4 <;> (try rw [two_smul]) <;> abel + ‖y‖ * ‖y‖ + ‖y - 2 • z‖ * ‖y - 2 • z‖ = 2 * (‖y - z‖ * ‖y - z‖ + ‖z‖ * ‖z‖) := by + convert parallelogram_identity (y - z) z using 4 <;> abel -private theorem add_left_aux4' (y z : E) : - (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 = - ‖y + z‖ * ‖y + z‖ - ‖y - z‖ * ‖y - z‖ := by - rw [add_left_aux3, add_left_aux4]; ring +variable (𝕜) private theorem add_left_aux5 (x y z : E) : - ‖(I : 𝕜) • (x + y) + z‖ * ‖(I : 𝕜) • (x + y) + z‖ = - (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + - ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖) / 2 - - ‖(I : 𝕜) • x - z‖ * ‖(I : 𝕜) • x - z‖ := by - rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm] - have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) + z) ((I : 𝕜) • x - z) - convert h₀ using 4 <;> · try simp only [two_smul, smul_add]; abel + ‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖ + = 2 * (‖(I : 𝕜) • (x + y) + z‖ * ‖(I : 𝕜) • (x + y) + z‖ + + ‖(I : 𝕜) • x - z‖ * ‖(I : 𝕜) • x - z‖) := by + convert parallelogram_identity ((I : 𝕜) • (x + y) + z) ((I : 𝕜) • x - z) using 4 <;> module private theorem add_left_aux6 (x y z : E) : - ‖(I : 𝕜) • (x + y) - z‖ * ‖(I : 𝕜) • (x + y) - z‖ = (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + - ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖) / 2 - - ‖(I : 𝕜) • x + z‖ * ‖(I : 𝕜) • x + z‖ := by - rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm] - have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) - z) ((I : 𝕜) • x + z) - convert h₀ using 4 <;> · try simp only [two_smul, smul_add]; abel + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖) + = 2 * (‖(I : 𝕜) • (x + y) - z‖ * ‖(I : 𝕜) • (x + y) - z‖ + + ‖(I : 𝕜) • x + z‖ * ‖(I : 𝕜) • x + z‖) := by + convert parallelogram_identity ((I : 𝕜) • (x + y) - z) ((I : 𝕜) • x + z) using 4 <;> module private theorem add_left_aux7 (y z : E) : - ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖ = - 2 * (‖(I : 𝕜) • y + z‖ * ‖(I : 𝕜) • y + z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := by - apply eq_sub_of_add_eq - have h₀ := parallelogram_identity ((I : 𝕜) • y + z) z - convert h₀ using 4 <;> · (try simp only [two_smul, smul_add]); abel + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖ + ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ = + 2 * (‖(I : 𝕜) • y + z‖ * ‖(I : 𝕜) • y + z‖ + ‖z‖ * ‖z‖) := by + convert parallelogram_identity ((I : 𝕜) • y + z) z using 4 <;> module private theorem add_left_aux8 (y z : E) : - ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖ = - 2 * (‖(I : 𝕜) • y - z‖ * ‖(I : 𝕜) • y - z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := by - apply eq_sub_of_add_eq' - have h₀ := parallelogram_identity ((I : 𝕜) • y - z) z - convert h₀ using 4 <;> · (try simp only [two_smul, smul_add]); abel + ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖ = + 2 * (‖(I : 𝕜) • y - z‖ * ‖(I : 𝕜) • y - z‖ + ‖z‖ * ‖z‖) := by + convert parallelogram_identity ((I : 𝕜) • y - z) z using 4 <;> module + +variable {𝕜} theorem add_left (x y z : E) : inner_ 𝕜 (x + y) z = inner_ 𝕜 x z + inner_ 𝕜 y z := by - simp only [inner_, ← mul_add] - congr - simp only [mul_assoc, ← map_mul, add_sub_assoc, ← mul_sub, ← map_sub] - rw [add_add_add_comm] - simp only [← map_add, ← mul_add] - congr - · rw [← add_sub_assoc, add_left_aux2', add_left_aux4'] - · rw [add_left_aux5, add_left_aux6, add_left_aux7, add_left_aux8] - simp only [map_sub, map_mul, map_add, div_eq_mul_inv] - ring - -theorem nat (n : ℕ) (x y : E) : inner_ 𝕜 ((n : 𝕜) • x) y = (n : 𝕜) * inner_ 𝕜 x y := by - induction' n with n ih - · simp only [inner_, zero_sub, Nat.cast_zero, zero_mul, - eq_self_iff_true, zero_smul, zero_add, mul_zero, sub_self, norm_neg, smul_zero] - · simp only [Nat.cast_succ, add_smul, one_smul] - rw [add_left, ih, add_mul, one_mul] - -private theorem nat_prop (r : ℕ) : innerProp' E (r : 𝕜) := fun x y => by - simp only [map_natCast]; exact nat r x y - -private theorem int_prop (n : ℤ) : innerProp' E (n : 𝕜) := by - intro x y - rw [← n.sign_mul_natAbs] - simp only [Int.cast_natCast, map_natCast, map_intCast, Int.cast_mul, map_mul, mul_smul] - obtain hn | rfl | hn := lt_trichotomy n 0 - · rw [Int.sign_eq_neg_one_of_neg hn, innerProp_neg_one ((n.natAbs : 𝕜) • x), nat] - simp only [map_neg, neg_mul, one_mul, mul_eq_mul_left_iff, Int.natAbs_eq_zero, - eq_self_iff_true, Int.cast_one, map_one, neg_inj, Nat.cast_eq_zero, Int.cast_neg] - · simp only [inner_, Int.cast_zero, zero_sub, Nat.cast_zero, zero_mul, - eq_self_iff_true, Int.sign_zero, zero_smul, zero_add, mul_zero, smul_zero, - sub_self, norm_neg, Int.natAbs_zero] - · rw [Int.sign_eq_one_of_pos hn] - simp only [one_mul, mul_eq_mul_left_iff, Int.natAbs_eq_zero, eq_self_iff_true, - Int.cast_one, one_smul, Nat.cast_eq_zero, nat] + have H_re := congr(- $(add_left_aux1 x y z) + $(add_left_aux2 x y z) + + $(add_left_aux3 y z) - $(add_left_aux4 y z)) + have H_im := congr(- $(add_left_aux5 𝕜 x y z) + $(add_left_aux6 𝕜 x y z) + + $(add_left_aux7 𝕜 y z) - $(add_left_aux8 𝕜 y z)) + have H := congr(𝓚 $H_re + I * 𝓚 $H_im) + simp only [inner_, map_add, map_sub, map_neg, map_mul, map_ofNat] at H ⊢ + linear_combination H / 8 private theorem rat_prop (r : ℚ) : innerProp' E (r : 𝕜) := by intro x y - have : (r.den : 𝕜) ≠ 0 := by - haveI : CharZero 𝕜 := RCLike.charZero_rclike - exact mod_cast r.pos.ne' - rw [← r.num_div_den, ← mul_right_inj' this, ← nat r.den _ y, smul_smul, Rat.cast_div] - simp only [map_natCast, Rat.cast_natCast, map_intCast, Rat.cast_intCast, map_div₀] - rw [← mul_assoc, mul_div_cancel₀ _ this, int_prop _ x, map_intCast] + let hom : 𝕜 →ₗ[ℚ] 𝕜 := AddMonoidHom.toRatLinearMap <| + AddMonoidHom.mk' (fun r ↦ inner_ 𝕜 (r • x) y) <| fun a b ↦ by + simpa [add_smul] using add_left (a • x) (b • x) y + simpa [hom, Rat.smul_def] using map_smul hom r 1 private theorem real_prop (r : ℝ) : innerProp' E (r : 𝕜) := by intro x y @@ -276,20 +202,15 @@ private theorem real_prop (r : ℝ) : innerProp' E (r : 𝕜) := by private theorem I_prop : innerProp' E (I : 𝕜) := by by_cases hI : (I : 𝕜) = 0 - · rw [hI, ← Nat.cast_zero]; exact nat_prop _ + · rw [hI] + simpa using real_prop (𝕜 := 𝕜) 0 intro x y - have hI' : (-I : 𝕜) * I = 1 := by rw [← inv_I, inv_mul_cancel₀ hI] - rw [conj_I, inner_, inner_, mul_left_comm] - congr 1 - rw [smul_smul, I_mul_I_of_nonzero hI, neg_one_smul] - rw [mul_sub, mul_add, mul_sub, mul_assoc I (𝓚 ‖I • x - y‖), ← mul_assoc (-I) I, hI', one_mul, - mul_assoc I (𝓚 ‖I • x + y‖), ← mul_assoc (-I) I, hI', one_mul] + have hI' := I_mul_I_of_nonzero hI + rw [conj_I, inner_, inner_, mul_left_comm, smul_smul, hI', neg_one_smul] have h₁ : ‖-x - y‖ = ‖x + y‖ := by rw [← neg_add', norm_neg] have h₂ : ‖-x + y‖ = ‖x - y‖ := by rw [← neg_sub, norm_neg, sub_eq_neg_add] rw [h₁, h₂] - simp only [sub_eq_add_neg, mul_assoc] - rw [← neg_mul_eq_neg_mul, ← neg_mul_eq_neg_mul] - abel + linear_combination (- 𝓚 ‖(I : 𝕜) • x - y‖ ^ 2 + 𝓚 ‖(I : 𝕜) • x + y‖ ^ 2) * hI' / 4 theorem innerProp (r : 𝕜) : innerProp' E r := by intro x y From 74620cae9fe1197241de26e09fa00c93f5cbb4b2 Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Wed, 2 Oct 2024 18:00:51 +0000 Subject: [PATCH 131/174] feat(ModelTheory/Equivalence): Implication of formulas modulo a theory (#16800) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Defines `FirstOrder.Language.Theory.Implies`: `φ ⟹[T] ψ` indicates that `φ` implies `ψ` in models of `T`. Relates `Theory.Implies` to `Theory.SemanticallyEquivalent`. Proves lemmas that will set up a `BooleanAlgebra` instance on formulas modulo a theory Co-authored-by: Yaël Dillies Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> --- Mathlib/ModelTheory/Complexity.lean | 10 +- Mathlib/ModelTheory/Equivalence.lean | 162 +++++++++++++++++++++------ 2 files changed, 134 insertions(+), 38 deletions(-) diff --git a/Mathlib/ModelTheory/Complexity.lean b/Mathlib/ModelTheory/Complexity.lean index f6af8bfbb9ccb..e4d9bfc7ef31f 100644 --- a/Mathlib/ModelTheory/Complexity.lean +++ b/Mathlib/ModelTheory/Complexity.lean @@ -291,7 +291,7 @@ theorem IsQF.induction_on_sup_not {P : L.BoundedFormula α n → Prop} {φ : L.B ∀ {φ₁ φ₂ : L.BoundedFormula α n}, (φ₁ ⇔[∅] φ₂) → (P φ₁ ↔ P φ₂)) : P φ := IsQF.recOn h hf @(ha) fun {φ₁ φ₂} _ _ h1 h2 => - (hse (φ₁.imp_semanticallyEquivalent_not_sup φ₂)).2 (hsup (hnot h1) h2) + (hse (φ₁.imp_iff_not_sup φ₂)).2 (hsup (hnot h1) h2) theorem IsQF.induction_on_inf_not {P : L.BoundedFormula α n → Prop} {φ : L.BoundedFormula α n} (h : IsQF φ) (hf : P (⊥ : L.BoundedFormula α n)) @@ -302,10 +302,10 @@ theorem IsQF.induction_on_inf_not {P : L.BoundedFormula α n → Prop} {φ : L.B P φ := h.induction_on_sup_not hf ha (fun {φ₁ φ₂} h1 h2 => - (hse (φ₁.sup_semanticallyEquivalent_not_inf_not φ₂)).2 (hnot (hinf (hnot h1) (hnot h2)))) + (hse (φ₁.sup_iff_not_inf_not φ₂)).2 (hnot (hinf (hnot h1) (hnot h2)))) (fun {_} => hnot) fun {_ _} => hse -theorem semanticallyEquivalent_toPrenex (φ : L.BoundedFormula α n) : +theorem iff_toPrenex (φ : L.BoundedFormula α n) : φ ⇔[∅] φ.toPrenex := fun M v xs => by rw [realize_iff, realize_toPrenex] @@ -317,7 +317,7 @@ theorem induction_on_all_ex {P : ∀ {m}, L.BoundedFormula α m → Prop} (φ : (φ₁ ⇔[∅] φ₂) → (P φ₁ ↔ P φ₂)) : P φ := by suffices h' : ∀ {m} {φ : L.BoundedFormula α m}, φ.IsPrenex → P φ from - (hse φ.semanticallyEquivalent_toPrenex).2 (h' φ.toPrenex_isPrenex) + (hse φ.iff_toPrenex).2 (h' φ.toPrenex_isPrenex) intro m φ hφ induction hφ with | of_isQF hφ => exact hqf hφ @@ -332,7 +332,7 @@ theorem induction_on_exists_not {P : ∀ {m}, L.BoundedFormula α m → Prop} ( (φ₁ ⇔[∅] φ₂) → (P φ₁ ↔ P φ₂)) : P φ := φ.induction_on_all_ex (fun {_ _} => hqf) - (fun {_ φ} hφ => (hse φ.all_semanticallyEquivalent_not_ex_not).2 (hnot (hex (hnot hφ)))) + (fun {_ φ} hφ => (hse φ.all_iff_not_ex_not).2 (hnot (hex (hnot hφ)))) (fun {_ _} => hex) fun {_ _ _} => hse /-- A universal formula is a formula defined by applying only universal quantifiers to a diff --git a/Mathlib/ModelTheory/Equivalence.lean b/Mathlib/ModelTheory/Equivalence.lean index e824467c1e629..42f8aff5f8d70 100644 --- a/Mathlib/ModelTheory/Equivalence.lean +++ b/Mathlib/ModelTheory/Equivalence.lean @@ -9,13 +9,12 @@ import Mathlib.ModelTheory.Satisfiability # Equivalence of Formulas ## Main Definitions -- `FirstOrder.Language.Theory.SemanticallyEquivalent`: `φ ⇔[T] ψ` indicates that `φ` and `ψ` are - equivalent formulas or sentences in models of `T`. +- `FirstOrder.Language.Theory.Imp`: `φ ⟹[T] ψ` indicates that `φ` implies `ψ` in models of `T`. +- `FirstOrder.Language.Theory.Iff`: `φ ⇔[T] ψ` indicates that `φ` and `ψ` are equivalent formulas or + sentences in models of `T`. ## TODO -- Add a definition of implication modulo a theory `T`, with `φ ⇒[T] ψ` defined analogously to - `φ ⇔[T] ψ`. -- Construct the quotient of `L.Formula α` modulo `⇔[T]` and its Boolean Algebra structure. +- Define the quotient of `L.Formula α` modulo `⇔[T]` and its Boolean Algebra structure. -/ @@ -34,24 +33,107 @@ variable {M : Type*} [Nonempty M] [L.Structure M] [M ⊨ T] namespace Theory +/-- `φ ⟹[T] ψ` indicates that `φ` implies `ψ` in models of `T`. -/ +protected def Imp (T : L.Theory) (φ ψ : L.BoundedFormula α n) : Prop := + T ⊨ᵇ φ.imp ψ + +@[inherit_doc FirstOrder.Language.Theory.Imp] +scoped[FirstOrder] notation:51 φ:50 " ⟹[" T "] " ψ:51 => Language.Theory.Imp T φ ψ + +namespace Imp + +@[refl] +protected theorem refl (φ : L.BoundedFormula α n) : φ ⟹[T] φ := fun _ _ _ => id + +instance : IsRefl (L.BoundedFormula α n) T.Imp := ⟨Imp.refl⟩ + +@[trans] +protected theorem trans {φ ψ θ : L.BoundedFormula α n} (h1 : φ ⟹[T] ψ) (h2 : ψ ⟹[T] θ) : + φ ⟹[T] θ := fun M v xs => (h2 M v xs) ∘ (h1 M v xs) + +instance : IsTrans (L.BoundedFormula α n) T.Imp := ⟨fun _ _ _ => Imp.trans⟩ + +end Imp + +section Imp + +lemma bot_imp (φ : L.BoundedFormula α n) : ⊥ ⟹[T] φ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_bot, false_implies] + +lemma imp_top (φ : L.BoundedFormula α n) : φ ⟹[T] ⊤ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_top, implies_true] + +lemma imp_sup_left (φ ψ : L.BoundedFormula α n) : φ ⟹[T] φ ⊔ ψ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_sup] + exact Or.inl + +lemma imp_sup_right (φ ψ : L.BoundedFormula α n) : ψ ⟹[T] φ ⊔ ψ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_sup] + exact Or.inr + +lemma sup_imp {φ ψ θ : L.BoundedFormula α n} (h₁ : φ ⟹[T] θ) (h₂ : ψ ⟹[T] θ) : + φ ⊔ ψ ⟹[T] θ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_sup] + exact fun h => h.elim (h₁ M v xs) (h₂ M v xs) + +lemma sup_imp_iff {φ ψ θ : L.BoundedFormula α n} : + (φ ⊔ ψ ⟹[T] θ) ↔ (φ ⟹[T] θ) ∧ (ψ ⟹[T] θ) := + ⟨fun h => ⟨(imp_sup_left _ _).trans h, (imp_sup_right _ _).trans h⟩, + fun ⟨h₁, h₂⟩ => sup_imp h₁ h₂⟩ + +lemma inf_imp_left (φ ψ : L.BoundedFormula α n) : φ ⊓ ψ ⟹[T] φ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_inf] + exact And.left + +lemma inf_imp_right (φ ψ : L.BoundedFormula α n) : φ ⊓ ψ ⟹[T] ψ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_inf] + exact And.right + +lemma imp_inf {φ ψ θ : L.BoundedFormula α n} (h₁ : φ ⟹[T] ψ) (h₂ : φ ⟹[T] θ) : + φ ⟹[T] ψ ⊓ θ := fun M v xs => by + simp only [BoundedFormula.realize_imp, BoundedFormula.realize_inf] + exact fun h => ⟨h₁ M v xs h, h₂ M v xs h⟩ + +lemma imp_inf_iff {φ ψ θ : L.BoundedFormula α n} : + (φ ⟹[T] ψ ⊓ θ) ↔ (φ ⟹[T] ψ) ∧ (φ ⟹[T] θ) := + ⟨fun h => ⟨h.trans (inf_imp_left _ _), h.trans (inf_imp_right _ _)⟩, + fun ⟨h₁, h₂⟩ => imp_inf h₁ h₂⟩ + +end Imp + /-- Two (bounded) formulas are semantically equivalent over a theory `T` when they have the same interpretation in every model of `T`. (This is also known as logical equivalence, which also has a proof-theoretic definition.) -/ -def SemanticallyEquivalent (T : L.Theory) (φ ψ : L.BoundedFormula α n) : Prop := +protected def Iff (T : L.Theory) (φ ψ : L.BoundedFormula α n) : Prop := T ⊨ᵇ φ.iff ψ -@[inherit_doc FirstOrder.Language.Theory.SemanticallyEquivalent] -scoped[FirstOrder] notation:25 φ " ⇔[" T "] " ψ => Language.Theory.SemanticallyEquivalent T φ ψ +@[inherit_doc FirstOrder.Language.Theory.Iff] +scoped[FirstOrder] +notation:51 φ:50 " ⇔[" T "] " ψ:51 => Language.Theory.Iff T φ ψ +theorem iff_iff_imp_and_imp {φ ψ : L.BoundedFormula α n} : + (φ ⇔[T] ψ) ↔ (φ ⟹[T] ψ) ∧ (ψ ⟹[T] φ) := by + simp only [Theory.Imp, ModelsBoundedFormula, BoundedFormula.realize_imp, ← forall_and, + Theory.Iff, BoundedFormula.realize_iff, iff_iff_implies_and_implies] -namespace SemanticallyEquivalent +theorem imp_antisymm {φ ψ : L.BoundedFormula α n} (h₁ : φ ⟹[T] ψ) (h₂ : ψ ⟹[T] φ) : + φ ⇔[T] ψ := + iff_iff_imp_and_imp.2 ⟨h₁, h₂⟩ + +namespace Iff + +protected theorem mp {φ ψ : L.BoundedFormula α n} (h : φ ⇔[T] ψ) : + φ ⟹[T] ψ := (iff_iff_imp_and_imp.1 h).1 + +protected theorem mpr {φ ψ : L.BoundedFormula α n} (h : φ ⇔[T] ψ) : + ψ ⟹[T] φ := (iff_iff_imp_and_imp.1 h).2 @[refl] protected theorem refl (φ : L.BoundedFormula α n) : φ ⇔[T] φ := fun M v xs => by rw [BoundedFormula.realize_iff] -instance : IsRefl (L.BoundedFormula α n) T.SemanticallyEquivalent := - ⟨SemanticallyEquivalent.refl⟩ +instance : IsRefl (L.BoundedFormula α n) T.Iff := + ⟨Iff.refl⟩ @[symm] protected theorem symm {φ ψ : L.BoundedFormula α n} @@ -59,6 +141,9 @@ protected theorem symm {φ ψ : L.BoundedFormula α n} rw [BoundedFormula.realize_iff, Iff.comm, ← BoundedFormula.realize_iff] exact h M v xs +instance : IsSymm (L.BoundedFormula α n) T.Iff := + ⟨fun _ _ => Iff.symm⟩ + @[trans] protected theorem trans {φ ψ θ : L.BoundedFormula α n} (h1 : φ ⇔[T] ψ) (h2 : ψ ⇔[T] θ) : @@ -68,6 +153,9 @@ protected theorem trans {φ ψ θ : L.BoundedFormula α n} rw [BoundedFormula.realize_iff] at * exact ⟨h2'.1 ∘ h1'.1, h1'.2 ∘ h2'.2⟩ +instance : IsTrans (L.BoundedFormula α n) T.Iff := + ⟨fun _ _ _ => Iff.trans⟩ + theorem realize_bd_iff {φ ψ : L.BoundedFormula α n} (h : φ ⇔[T] ψ) {v : α → M} {xs : Fin n → M} : φ.Realize v xs ↔ ψ.Realize v xs := BoundedFormula.realize_iff.1 (h.realize_boundedFormula M) @@ -84,33 +172,33 @@ theorem models_sentence_iff {φ ψ : L.Sentence} {M : Type*} [Nonempty M] protected theorem all {φ ψ : L.BoundedFormula α (n + 1)} (h : φ ⇔[T] ψ) : φ.all ⇔[T] ψ.all := by - simp_rw [SemanticallyEquivalent, ModelsBoundedFormula, BoundedFormula.realize_iff, + simp_rw [Theory.Iff, ModelsBoundedFormula, BoundedFormula.realize_iff, BoundedFormula.realize_all] exact fun M v xs => forall_congr' fun a => h.realize_bd_iff protected theorem ex {φ ψ : L.BoundedFormula α (n + 1)} (h : φ ⇔[T] ψ) : φ.ex ⇔[T] ψ.ex := by - simp_rw [SemanticallyEquivalent, ModelsBoundedFormula, BoundedFormula.realize_iff, + simp_rw [Theory.Iff, ModelsBoundedFormula, BoundedFormula.realize_iff, BoundedFormula.realize_ex] exact fun M v xs => exists_congr fun a => h.realize_bd_iff protected theorem not {φ ψ : L.BoundedFormula α n} (h : φ ⇔[T] ψ) : φ.not ⇔[T] ψ.not := by - simp_rw [SemanticallyEquivalent, ModelsBoundedFormula, BoundedFormula.realize_iff, + simp_rw [Theory.Iff, ModelsBoundedFormula, BoundedFormula.realize_iff, BoundedFormula.realize_not] exact fun M v xs => not_congr h.realize_bd_iff protected theorem imp {φ ψ φ' ψ' : L.BoundedFormula α n} (h : φ ⇔[T] ψ) (h' : φ' ⇔[T] ψ') : (φ.imp φ') ⇔[T] (ψ.imp ψ') := by - simp_rw [SemanticallyEquivalent, ModelsBoundedFormula, BoundedFormula.realize_iff, + simp_rw [Theory.Iff, ModelsBoundedFormula, BoundedFormula.realize_iff, BoundedFormula.realize_imp] exact fun M v xs => imp_congr h.realize_bd_iff h'.realize_bd_iff -end SemanticallyEquivalent +end Iff /-- Semantic equivalence forms an equivalence relation on formulas. -/ -def semanticallyEquivalentSetoid (T : L.Theory) : Setoid (L.BoundedFormula α n) where - r := SemanticallyEquivalent T +def iffSetoid (T : L.Theory) : Setoid (L.BoundedFormula α n) where + r := T.Iff iseqv := ⟨fun _ => refl _, fun {_ _} h => h.symm, fun {_ _ _} h1 h2 => h1.trans h2⟩ end Theory @@ -119,45 +207,53 @@ namespace BoundedFormula variable (φ ψ : L.BoundedFormula α n) -theorem semanticallyEquivalent_not_not : φ ⇔[T] φ.not.not := fun M v xs => by +theorem iff_not_not : φ ⇔[T] φ.not.not := fun M v xs => by simp -theorem imp_semanticallyEquivalent_not_sup : (φ.imp ψ) ⇔[T] (φ.not ⊔ ψ) := +theorem imp_iff_not_sup : (φ.imp ψ) ⇔[T] (φ.not ⊔ ψ) := fun M v xs => by simp [imp_iff_not_or] -theorem sup_semanticallyEquivalent_not_inf_not : (φ ⊔ ψ) ⇔[T] (φ.not ⊓ ψ.not).not := +theorem sup_iff_not_inf_not : (φ ⊔ ψ) ⇔[T] (φ.not ⊓ ψ.not).not := fun M v xs => by simp [imp_iff_not_or] -theorem inf_semanticallyEquivalent_not_sup_not : (φ ⊓ ψ) ⇔[T] (φ.not ⊔ ψ.not).not := +theorem inf_iff_not_sup_not : (φ ⊓ ψ) ⇔[T] (φ.not ⊔ ψ.not).not := fun M v xs => by simp -theorem all_semanticallyEquivalent_not_ex_not (φ : L.BoundedFormula α (n + 1)) : +theorem all_iff_not_ex_not (φ : L.BoundedFormula α (n + 1)) : φ.all ⇔[T] φ.not.ex.not := fun M v xs => by simp -theorem ex_semanticallyEquivalent_not_all_not (φ : L.BoundedFormula α (n + 1)) : +theorem ex_iff_not_all_not (φ : L.BoundedFormula α (n + 1)) : φ.ex ⇔[T] φ.not.all.not := fun M v xs => by simp -theorem semanticallyEquivalent_all_liftAt : φ ⇔[T] (φ.liftAt 1 n).all := +theorem iff_all_liftAt : φ ⇔[T] (φ.liftAt 1 n).all := fun M v xs => by rw [realize_iff, realize_all_liftAt_one_self] +lemma inf_not_iff_bot : + φ ⊓ ∼φ ⇔[T] ⊥ := fun M v xs => by + simp only [realize_iff, realize_inf, realize_not, and_not_self, realize_bot] + +lemma sup_not_iff_top : + φ ⊔ ∼φ ⇔[T] ⊤ := fun M v xs => by + simp only [realize_iff, realize_sup, realize_not, realize_top, iff_true, or_not] + end BoundedFormula namespace Formula variable (φ ψ : L.Formula α) -theorem semanticallyEquivalent_not_not : φ ⇔[T] φ.not.not := - BoundedFormula.semanticallyEquivalent_not_not φ +theorem iff_not_not : φ ⇔[T] φ.not.not := + BoundedFormula.iff_not_not φ -theorem imp_semanticallyEquivalent_not_sup : (φ.imp ψ) ⇔[T] (φ.not ⊔ ψ) := - BoundedFormula.imp_semanticallyEquivalent_not_sup φ ψ +theorem imp_iff_not_sup : (φ.imp ψ) ⇔[T] (φ.not ⊔ ψ) := + BoundedFormula.imp_iff_not_sup φ ψ -theorem sup_semanticallyEquivalent_not_inf_not : (φ ⊔ ψ) ⇔[T] (φ.not ⊓ ψ.not).not := - BoundedFormula.sup_semanticallyEquivalent_not_inf_not φ ψ +theorem sup_iff_not_inf_not : (φ ⊔ ψ) ⇔[T] (φ.not ⊓ ψ.not).not := + BoundedFormula.sup_iff_not_inf_not φ ψ -theorem inf_semanticallyEquivalent_not_sup_not : (φ ⊓ ψ) ⇔[T] (φ.not ⊔ ψ.not).not := - BoundedFormula.inf_semanticallyEquivalent_not_sup_not φ ψ +theorem inf_iff_not_sup_not : (φ ⊓ ψ) ⇔[T] (φ.not ⊔ ψ.not).not := + BoundedFormula.inf_iff_not_sup_not φ ψ end Formula From f01032b3079b73ededb6ca35c5196964a6d537d0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 18:00:52 +0000 Subject: [PATCH 132/174] =?UTF-8?q?chore:=20`one=5Fle=5Fpow=5Fof=5Fone=5Fl?= =?UTF-8?q?e`=20->=20`one=5Fle=5Fpow=E2=82=80`=20(#17215)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rename: * `pow_le_one` -> `pow_le_one₀` * `pow_lt_one` -> `pow_lt_one₀` * `one_le_pow_of_one_le` -> `one_le_pow₀` * `one_lt_pow` -> `one_lt_pow₀` and make all non-order arguments implicit (those are not rewrite lemmas, and where the explicit arguments are located is basically random throughout the library) From LeanAPAP --- Archive/Imo/Imo2006Q5.lean | 2 +- Archive/Imo/Imo2013Q5.lean | 8 ++--- .../Algebra/Order/CauSeq/BigOperators.lean | 2 +- Mathlib/Algebra/Order/Field/Power.lean | 6 ++-- .../Algebra/Order/Interval/Set/Instances.lean | 4 +-- Mathlib/Algebra/Order/Ring/Basic.lean | 30 ++++++++++--------- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/Constructions.lean | 4 +-- Mathlib/Analysis/Analytic/RadiusLiminf.lean | 2 +- .../Analysis/Calculus/FDeriv/Measurable.lean | 4 +-- Mathlib/Analysis/Calculus/TangentCone.lean | 2 +- Mathlib/Analysis/Complex/AbelLimit.lean | 2 +- .../Analysis/Distribution/SchwartzSpace.lean | 6 ++-- .../Fourier/FourierTransformDeriv.lean | 2 +- .../Analysis/SpecialFunctions/Stirling.lean | 2 +- .../SpecialFunctions/Trigonometric/Basic.lean | 4 +-- Mathlib/Analysis/SpecificLimits/Basic.lean | 4 +-- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 11 ++++--- .../SimpleGraph/Regularity/Bound.lean | 4 +-- .../SimpleGraph/Regularity/Chunk.lean | 2 +- .../AkraBazzi/GrowsPolynomially.lean | 2 +- Mathlib/Data/Complex/Exponential.lean | 4 +-- Mathlib/Data/Real/Pi/Irrational.lean | 2 +- Mathlib/GroupTheory/Exponent.lean | 2 +- Mathlib/GroupTheory/FiniteAbelian.lean | 2 +- Mathlib/GroupTheory/OrderOfElement.lean | 4 +-- Mathlib/GroupTheory/PGroup.lean | 2 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 2 +- .../Group/GeometryOfNumbers.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 9 +++--- Mathlib/NumberTheory/FermatPsp.lean | 2 +- Mathlib/NumberTheory/Harmonic/Int.lean | 2 +- Mathlib/NumberTheory/Liouville/Basic.lean | 2 +- .../Liouville/LiouvilleNumber.lean | 2 +- Mathlib/NumberTheory/LucasLehmer.lean | 4 +-- .../CanonicalEmbedding/ConvexBody.lean | 6 ++-- .../NumberField/Discriminant.lean | 2 +- .../NumberTheory/NumberField/Embeddings.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/Probability/StrongLaw.lean | 4 +-- Mathlib/Topology/Algebra/Polynomial.lean | 3 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 2 +- 43 files changed, 84 insertions(+), 85 deletions(-) diff --git a/Archive/Imo/Imo2006Q5.lean b/Archive/Imo/Imo2006Q5.lean index 0e2398d7919ba..d317993a8c618 100644 --- a/Archive/Imo/Imo2006Q5.lean +++ b/Archive/Imo/Imo2006Q5.lean @@ -122,7 +122,7 @@ theorem Polynomial.iterate_comp_sub_X_ne {P : Polynomial ℤ} (hP : 1 < P.natDeg (hk : 0 < k) : P.comp^[k] X - X ≠ 0 := by rw [sub_ne_zero] apply_fun natDegree - simpa using (one_lt_pow hP hk.ne').ne' + simpa using (one_lt_pow₀ hP hk.ne').ne' /-- We solve the problem for the specific case k = 2 first. -/ theorem imo2006_q5' {P : Polynomial ℤ} (hP : 1 < P.natDegree) : diff --git a/Archive/Imo/Imo2013Q5.lean b/Archive/Imo/Imo2013Q5.lean index 48e7486e9c230..225e64ae35b02 100644 --- a/Archive/Imo/Imo2013Q5.lean +++ b/Archive/Imo/Imo2013Q5.lean @@ -38,9 +38,9 @@ theorem le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) have hterm : ∀ i : ℕ, i ∈ Finset.range n → 1 ≤ x ^ i * y ^ (n - 1 - i) := by intro i _ calc - 1 ≤ x ^ i := one_le_pow_of_one_le hx.le i + 1 ≤ x ^ i := one_le_pow₀ hx.le _ = x ^ i * 1 := by ring - _ ≤ x ^ i * y ^ (n - 1 - i) := by gcongr; apply one_le_pow_of_one_le hy.le + _ ≤ x ^ i * y ^ (n - 1 - i) := by gcongr; apply one_le_pow₀ hy.le calc (x - y) * (n : ℝ) = (n : ℝ) * (x - y) := by ring _ = (∑ _i ∈ Finset.range n, (1 : ℝ)) * (x - y) := by @@ -134,7 +134,7 @@ theorem fixed_point_of_pos_nat_pow {f : ℚ → ℝ} {n : ℕ} (hn : 0 < n) (H1 : ∀ x y, 0 < x → 0 < y → f (x * y) ≤ f x * f y) (H4 : ∀ n : ℕ, 0 < n → (n : ℝ) ≤ f n) (H5 : ∀ x : ℚ, 1 < x → (x : ℝ) ≤ f x) {a : ℚ} (ha1 : 1 < a) (hae : f a = a) : f (a ^ n) = a ^ n := by - have hh0 : (a : ℝ) ^ n ≤ f (a ^ n) := mod_cast H5 (a ^ n) (one_lt_pow ha1 hn.ne') + have hh0 : (a : ℝ) ^ n ≤ f (a ^ n) := mod_cast H5 (a ^ n) (one_lt_pow₀ ha1 hn.ne') have hh1 := calc f (a ^ n) ≤ f a ^ n := pow_f_le_f_pow hn ha1 H1 H4 @@ -206,7 +206,7 @@ theorem imo2013_q5 (f : ℚ → ℝ) (H1 : ∀ x y, 0 < x → 0 < y → f (x * y intro n hn calc (x : ℝ) ^ n - 1 < f (x ^ n) := - mod_cast fx_gt_xm1 (one_le_pow_of_one_le hx.le n) H1 H2 H4 + mod_cast fx_gt_xm1 (one_le_pow₀ hx.le) H1 H2 H4 _ ≤ f x ^ n := pow_f_le_f_pow hn hx H1 H4 have hx' : 1 < (x : ℝ) := mod_cast hx have hxp : 0 < x := by positivity diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 2ac291cf0c7a6..96a11d63d193a 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -188,7 +188,7 @@ lemma geo_series [Nontrivial β] (x : β) (hx1 : abv x < 1) : · gcongr exact sub_le_self _ (abv_pow abv x n ▸ abv_nonneg _ _) refine div_nonneg (sub_nonneg.2 ?_) (sub_nonneg.2 <| le_of_lt hx1) - exact pow_le_one _ (by positivity) hx1.le + exact pow_le_one₀ (by positivity) hx1.le · intro n _ rw [← one_mul (abv x ^ n), pow_succ'] gcongr diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 8dfd3676717c3..70c83a1869e07 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -31,7 +31,7 @@ theorem zpow_le_of_le (ha : 1 ≤ a) (h : m ≤ n) : a ^ m ≤ a ^ n := by calc a ^ m = a ^ m * 1 := (mul_one _).symm _ ≤ a ^ m * a ^ k := - mul_le_mul_of_nonneg_left (one_le_pow_of_one_le ha _) (zpow_nonneg ha₀.le _) + mul_le_mul_of_nonneg_left (one_le_pow₀ ha) (zpow_nonneg ha₀.le _) _ = a ^ n := by rw [← zpow_natCast, ← zpow_add₀ ha₀.ne', hk, add_sub_cancel] theorem zpow_le_one_of_nonpos (ha : 1 ≤ a) (hn : n ≤ 0) : a ^ n ≤ 1 := @@ -48,7 +48,7 @@ theorem Nat.zpow_ne_zero_of_pos {a : ℕ} (h : 0 < a) (n : ℤ) : (a : α) ^ n (Nat.zpow_pos_of_pos h n).ne' theorem one_lt_zpow (ha : 1 < a) : ∀ n : ℤ, 0 < n → 1 < a ^ n - | (n : ℕ), h => (zpow_natCast _ _).symm.subst (one_lt_pow ha <| Int.natCast_ne_zero.mp h.ne') + | (n : ℕ), h => (zpow_natCast _ _).symm.subst (one_lt_pow₀ ha <| Int.natCast_ne_zero.mp h.ne') | -[_+1], h => ((Int.negSucc_not_pos _).mp h).elim theorem zpow_strictMono (hx : 1 < a) : StrictMono (a ^ · : ℤ → α) := @@ -79,7 +79,7 @@ theorem zpow_le_iff_le (hx : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := @[simp] theorem div_pow_le (ha : 0 ≤ a) (hb : 1 ≤ b) (k : ℕ) : a / b ^ k ≤ a := - div_le_self ha <| one_le_pow_of_one_le hb _ + div_le_self ha <| one_le_pow₀ hb theorem zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : Injective (a ^ · : ℤ → α) := by rcases h₁.lt_or_lt with (H | H) diff --git a/Mathlib/Algebra/Order/Interval/Set/Instances.lean b/Mathlib/Algebra/Order/Interval/Set/Instances.lean index 74e41d229a3ff..da5a81838c6e1 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Instances.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Instances.lean @@ -101,7 +101,7 @@ instance mul : Mul (Icc (0 : α) 1) where mul p q := ⟨p * q, ⟨mul_nonneg p.2.1 q.2.1, mul_le_one p.2.2 q.2.1 q.2.2⟩⟩ instance pow : Pow (Icc (0 : α) 1) ℕ where - pow p n := ⟨p.1 ^ n, ⟨pow_nonneg p.2.1 n, pow_le_one n p.2.1 p.2.2⟩⟩ + pow p n := ⟨p.1 ^ n, ⟨pow_nonneg p.2.1 n, pow_le_one₀ p.2.1 p.2.2⟩⟩ @[simp, norm_cast] theorem coe_mul (x y : Icc (0 : α) 1) : ↑(x * y) = (x * y : α) := @@ -239,7 +239,7 @@ instance mul : Mul (Ioc (0 : α) 1) where mul p q := ⟨p.1 * q.1, ⟨mul_pos p.2.1 q.2.1, mul_le_one p.2.2 (le_of_lt q.2.1) q.2.2⟩⟩ instance pow : Pow (Ioc (0 : α) 1) ℕ where - pow p n := ⟨p.1 ^ n, ⟨pow_pos p.2.1 n, pow_le_one n (le_of_lt p.2.1) p.2.2⟩⟩ + pow p n := ⟨p.1 ^ n, ⟨pow_pos p.2.1 n, pow_le_one₀ (le_of_lt p.2.1) p.2.2⟩⟩ @[simp, norm_cast] theorem coe_mul (x y : Ioc (0 : α) 1) : ↑(x * y) = (x * y : α) := diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 0d45fc5b2b588..b7c4848e38016 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -61,23 +61,31 @@ theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y exact mul_le_mul_of_nonneg_left (ih (Nat.succ_ne_zero k)) h2 @[bound] -theorem pow_le_one : ∀ n : ℕ, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 +theorem pow_le_one₀ : ∀ {n : ℕ}, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 | 0, _, _ => (pow_zero a).le - | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one (pow_le_one n h₀ h₁) h₀ h₁) + | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one (pow_le_one₀ h₀ h₁) h₀ h₁) -theorem pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 +theorem pow_lt_one₀ (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → a ^ n < 1 | 0, h => (h rfl).elim | n + 1, _ => by rw [pow_succ'] - exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one _ h₀ h₁.le) + exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one₀ h₀ h₁.le) @[bound] -theorem one_le_pow_of_one_le (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n +theorem one_le_pow₀ (H : 1 ≤ a) : ∀ {n : ℕ}, 1 ≤ a ^ n | 0 => by rw [pow_zero] | n + 1 => by - rw [pow_succ'] - simpa only [mul_one] using - mul_le_mul H (one_le_pow_of_one_le H n) zero_le_one (le_trans zero_le_one H) + simpa only [pow_succ', mul_one] + using mul_le_mul H (one_le_pow₀ H) zero_le_one (zero_le_one.trans H) + +lemma one_lt_pow₀ (ha : 1 < a) : ∀ {n : ℕ}, n ≠ 0 → 1 < a ^ n + | 0, h => (h rfl).elim + | n + 1, _ => by rw [pow_succ']; exact one_lt_mul_of_lt_of_le ha (one_le_pow₀ ha.le) + +@[deprecated (since := "2024-09-28")] alias pow_le_one := pow_le_one₀ +@[deprecated (since := "2024-09-28")] alias pow_lt_one := pow_lt_one₀ +@[deprecated (since := "2024-09-28")] alias one_le_pow_of_one_le := one_le_pow₀ +@[deprecated (since := "2024-09-28")] alias one_lt_pow := one_lt_pow₀ theorem pow_right_mono (h : 1 ≤ a) : Monotone (a ^ ·) := monotone_nat_of_le_succ fun n => by @@ -101,12 +109,6 @@ theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n | n + 1 => by simpa only [pow_succ'] using mul_le_mul hab (pow_le_pow_left ha hab _) (pow_nonneg ha _) (ha.trans hab) -theorem one_lt_pow (ha : 1 < a) : ∀ {n : ℕ} (_ : n ≠ 0), 1 < a ^ n - | 0, h => (h rfl).elim - | n + 1, _ => by - rw [pow_succ'] - exact one_lt_mul_of_lt_of_le ha (one_le_pow_of_one_le ha.le _) - lemma pow_add_pow_le' (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ n + b ^ n ≤ 2 * (a + b) ^ n := by rw [two_mul] exact add_le_add (pow_le_pow_left ha (le_add_of_nonneg_right hb) _) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 19befbc09fbe9..399762603a605 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -225,7 +225,7 @@ theorem lt_radius_of_isBigO (h₀ : r ≠ 0) {a : ℝ} (ha : a ∈ Ioo (-1 : ℝ theorem norm_mul_pow_le_of_lt_radius (p : FormalMultilinearSeries 𝕜 E F) {r : ℝ≥0} (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ‖p n‖ * (r : ℝ) ^ n ≤ C := let ⟨_, ha, C, hC, h⟩ := p.norm_mul_pow_le_mul_pow_of_lt_radius h - ⟨C, hC, fun n => (h n).trans <| mul_le_of_le_one_right hC.lt.le (pow_le_one _ ha.1.le ha.2.le)⟩ + ⟨C, hC, fun n => (h n).trans <| mul_le_of_le_one_right hC.lt.le (pow_le_one₀ ha.1.le ha.2.le)⟩ /-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` is bounded. -/ theorem norm_le_div_pow_of_pos_of_lt_radius (p : FormalMultilinearSeries 𝕜 E F) {r : ℝ≥0} diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 5b8b4b639daa3..42de038a5780f 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -734,7 +734,7 @@ lemma one_le_formalMultilinearSeries_geometric_radius (𝕜 : Type*) [Nontrivial apply le_trans ?_ (formalMultilinearSeries_geometric_apply_norm_le 𝕜 A n) conv_rhs => rw [← mul_one (‖formalMultilinearSeries_geometric 𝕜 A n‖)] gcongr - exact pow_le_one _ (coe_nonneg r) hr.le + exact pow_le_one₀ (coe_nonneg r) hr.le lemma formalMultilinearSeries_geometric_radius (𝕜 : Type*) [NontriviallyNormedField 𝕜] (A : Type*) [NormedRing A] [NormOneClass A] [NormedAlgebra 𝕜 A] : @@ -760,7 +760,7 @@ lemma formalMultilinearSeries_geometric_radius (𝕜 : Type*) [NontriviallyNorme simp_rw [formalMultilinearSeries_geometric_apply_norm, one_mul] refine isBigO_of_le atTop (fun n ↦ ?_) rw [norm_one, Real.norm_of_nonneg (pow_nonneg (coe_nonneg r) _)] - exact pow_le_one _ (coe_nonneg r) hr.le + exact pow_le_one₀ (coe_nonneg r) hr.le lemma hasFPowerSeriesOnBall_inverse_one_sub (𝕜 : Type*) [NontriviallyNormedField 𝕜] diff --git a/Mathlib/Analysis/Analytic/RadiusLiminf.lean b/Mathlib/Analysis/Analytic/RadiusLiminf.lean index 175d6acca7461..6aefcab0e900e 100644 --- a/Mathlib/Analysis/Analytic/RadiusLiminf.lean +++ b/Mathlib/Analysis/Analytic/RadiusLiminf.lean @@ -53,7 +53,7 @@ theorem radius_eq_liminf : refine H.mp ((eventually_gt_atTop 0).mono fun n hn₀ hn => (this _ hn₀).2 (NNReal.coe_le_coe.1 ?_)) push_cast - exact (le_abs_self _).trans (hn.trans (pow_le_one _ ha.1.le ha.2.le)) + exact (le_abs_self _).trans (hn.trans (pow_le_one₀ ha.1.le ha.2.le)) · refine p.le_radius_of_isBigO (IsBigO.of_bound 1 ?_) refine (eventually_lt_of_lt_liminf hr).mp ((eventually_gt_atTop 0).mono fun n hn₀ hn => ?_) simpa using NNReal.coe_le_coe.2 ((this _ hn₀).1 hn.le) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 3066b3ab3c333..28e1539f0e516 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -303,7 +303,7 @@ theorem D_subset_differentiable_set {K : Set (E →L[𝕜] F)} (hK : IsComplete · simp [y_pos] have yzero : 0 < ‖y‖ := norm_pos_iff.mpr y_pos have y_lt : ‖y‖ < (1 / 2) ^ (n e + 1) := by simpa using mem_ball_iff_norm.1 hy - have yone : ‖y‖ ≤ 1 := le_trans y_lt.le (pow_le_one _ (by norm_num) (by norm_num)) + have yone : ‖y‖ ≤ 1 := le_trans y_lt.le (pow_le_one₀ (by norm_num) (by norm_num)) -- define the scale `k`. obtain ⟨k, hk, h'k⟩ : ∃ k : ℕ, (1 / 2) ^ (k + 1) < ‖y‖ ∧ ‖y‖ ≤ (1 / 2) ^ k := exists_nat_pow_near_of_lt_one yzero yone (by norm_num : (0 : ℝ) < 1 / 2) @@ -638,7 +638,7 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : · simp only [sub_self, zero_smul, norm_zero, mul_zero, le_rfl] have yzero : 0 < y - x := sub_pos.2 xy have y_le : y - x ≤ (1 / 2) ^ (n e + 1) := by linarith [hy.2] - have yone : y - x ≤ 1 := le_trans y_le (pow_le_one _ (by norm_num) (by norm_num)) + have yone : y - x ≤ 1 := le_trans y_le (pow_le_one₀ (by norm_num) (by norm_num)) -- define the scale `k`. obtain ⟨k, hk, h'k⟩ : ∃ k : ℕ, (1 / 2) ^ (k + 1) < y - x ∧ y - x ≤ (1 / 2) ^ k := exists_nat_pow_near_of_lt_one yzero yone (by norm_num : (0 : ℝ) < 1 / 2) diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index 280e5356cb943..53b15bc072b28 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -196,7 +196,7 @@ theorem mem_tangentCone_of_openSegment_subset {s : Set G} {x y : G} (h : openSeg rw [openSegment_eq_image] refine ⟨(1 / 2) ^ n, ⟨?_, ?_⟩, ?_⟩ · exact pow_pos one_half_pos _ - · exact pow_lt_one one_half_pos.le one_half_lt_one hn + · exact pow_lt_one₀ one_half_pos.le one_half_lt_one hn · simp only [sub_smul, one_smul, smul_sub]; abel /-- If a subset of a real vector space contains a segment, then the direction of this diff --git a/Mathlib/Analysis/Complex/AbelLimit.lean b/Mathlib/Analysis/Complex/AbelLimit.lean index d00d00856facc..84347303b29c2 100644 --- a/Mathlib/Analysis/Complex/AbelLimit.lean +++ b/Mathlib/Analysis/Complex/AbelLimit.lean @@ -206,7 +206,7 @@ theorem tendsto_tsum_powerSeries_nhdsWithin_stolzSet calc _ ≤ ‖1 - z‖ * ∑ i ∈ range B₁, ‖l - s (i + 1)‖ := by gcongr; nth_rw 3 [← mul_one ‖_‖] - gcongr; exact pow_le_one _ (norm_nonneg _) zn.le + gcongr; exact pow_le_one₀ (norm_nonneg _) zn.le _ ≤ ‖1 - z‖ * (F + 1) := by gcongr; linarith only _ < _ := by rwa [norm_sub_rev, lt_div_iff (by positivity)] at zd have S₂ : ‖1 - z‖ * ∑ i ∈ Ico B₁ (max B₁ B₂), ‖l - s (i + 1)‖ * ‖z‖ ^ i < ε / 4 := diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 310970a55e4ef..43f9326f818a8 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -616,7 +616,7 @@ lemma pow_mul_le_of_le_of_pow_mul_le {C₁ C₂ : ℝ} {k l : ℕ} {x f : ℝ} ( rw [this] rcases le_total x 1 with h'x|h'x · gcongr - · apply (pow_le_one k hx h'x).trans + · apply (pow_le_one₀ hx h'x).trans apply Real.one_le_rpow_of_pos_of_le_one_of_nonpos · linarith · linarith @@ -854,7 +854,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) refine add_le_add ?_ (hg_upper' x) nth_rw 1 [← one_mul (1 : ℝ)] gcongr - apply one_le_pow_of_one_le + apply one_le_pow₀ simp only [le_add_iff_nonneg_right, norm_nonneg] have hbound : ∀ i, i ≤ n → ‖iteratedFDeriv ℝ i f (g x)‖ ≤ 2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k' := by @@ -871,7 +871,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' gcongr · exact le_trans (by simp [hC]) (le_self_pow (by simp [hC]) hN₁') - · refine le_self_pow (one_le_pow_of_one_le ?_ l) hN₁' + · refine le_self_pow (one_le_pow₀ ?_) hN₁' simp only [le_add_iff_nonneg_right, norm_nonneg] have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index e48a2caf91986..16957cece15d7 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -627,7 +627,7 @@ lemma pow_mul_norm_iteratedFDeriv_fourierIntegral_le [FiniteDimensional ℝ V] _ ≤ (2 * π) ^ n * (|L v w| ^ n * ‖iteratedFDeriv ℝ k (fourierIntegral 𝐞 μ L.toLinearMap₂ f) w‖) := by apply le_mul_of_one_le_left (by positivity) - apply one_le_pow_of_one_le + apply one_le_pow₀ linarith [one_le_pi_div_two] _ = ‖fourierPowSMulRight (-L.flip) (iteratedFDeriv ℝ k (fourierIntegral 𝐞 μ L.toLinearMap₂ f)) w n (fun _ ↦ v)‖ := by diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 0ff3d40698098..a1b4c1c5dd12d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -102,7 +102,7 @@ theorem log_stirlingSeq_diff_le_geo_sum (n : ℕ) : · simp_rw [← _root_.pow_succ'] at this exact this rw [one_div, inv_pow] - exact inv_lt_one (one_lt_pow ((lt_add_iff_pos_left 1).mpr <| by positivity) two_ne_zero) + exact inv_lt_one (one_lt_pow₀ ((lt_add_iff_pos_left 1).mpr <| by positivity) two_ne_zero) have hab (k : ℕ) : (1 : ℝ) / (2 * ↑(k + 1) + 1) * ((1 / (2 * ↑(n + 1) + 1)) ^ 2) ^ ↑(k + 1) ≤ (((1 : ℝ) / (2 * ↑(n + 1) + 1)) ^ 2) ^ ↑(k + 1) := by refine mul_le_of_le_one_left (pow_nonneg h_nonneg ↑(k + 1)) ?_ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index c31a03b44e31f..dfd9f1a9f4692 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -661,7 +661,7 @@ theorem sqrtTwoAddSeries_monotone_left {x y : ℝ} (h : x ≤ y) : theorem cos_pi_over_two_pow : ∀ n : ℕ, cos (π / 2 ^ (n + 1)) = sqrtTwoAddSeries 0 n / 2 | 0 => by simp | n + 1 => by - have A : (1 : ℝ) < 2 ^ (n + 1) := one_lt_pow one_lt_two n.succ_ne_zero + have A : (1 : ℝ) < 2 ^ (n + 1) := one_lt_pow₀ one_lt_two n.succ_ne_zero have B : π / 2 ^ (n + 1) < π := div_lt_self pi_pos A have C : 0 < π / 2 ^ (n + 1) := by positivity rw [pow_succ, div_mul_eq_div_div, cos_half, cos_pi_over_two_pow n, sqrtTwoAddSeries, @@ -690,7 +690,7 @@ theorem sin_pi_over_two_pow_succ (n : ℕ) : exact (sqrtTwoAddSeries_lt_two _).le refine mul_nonneg (sin_nonneg_of_nonneg_of_le_pi ?_ ?_) zero_le_two · positivity - · exact div_le_self pi_pos.le <| one_le_pow_of_one_le one_le_two _ + · exact div_le_self pi_pos.le <| one_le_pow₀ one_le_two @[simp] theorem cos_pi_div_four : cos (π / 4) = √2 / 2 := by diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 661e426e3784c..370cf43a6f16a 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -252,7 +252,7 @@ protected theorem ENNReal.tendsto_pow_atTop_nhds_top_iff {r : ℝ≥0∞} : specialize h_tends (Ioi_mem_nhds one_lt_top) simp only [Filter.mem_map, mem_atTop_sets, ge_iff_le, Set.mem_preimage, Set.mem_Ioi] at h_tends obtain ⟨n, hn⟩ := h_tends - exact lt_irrefl _ <| lt_of_lt_of_le (hn n le_rfl) <| pow_le_one n (zero_le _) r_le_one + exact lt_irrefl _ <| lt_of_lt_of_le (hn n le_rfl) <| pow_le_one₀ (zero_le _) r_le_one · intro r_gt_one have obs := @Tendsto.inv ℝ≥0∞ ℕ _ _ _ (fun n ↦ (r⁻¹)^n) atTop 0 simp only [ENNReal.tendsto_pow_atTop_nhds_zero_iff, inv_zero] at obs @@ -360,7 +360,7 @@ theorem ENNReal.tsum_geometric (r : ℝ≥0∞) : ∑' n : ℕ, r ^ n = (1 - r) (ENNReal.exists_nat_gt (lt_top_iff_ne_top.1 ha)).imp fun n hn ↦ lt_of_lt_of_le hn ?_ calc (n : ℝ≥0∞) = ∑ i ∈ range n, 1 := by rw [sum_const, nsmul_one, card_range] - _ ≤ ∑ i ∈ range n, r ^ i := by gcongr; apply one_le_pow_of_one_le' hr + _ ≤ ∑ i ∈ range n, r ^ i := by gcongr; apply one_le_pow₀ hr theorem ENNReal.tsum_geometric_add_one (r : ℝ≥0∞) : ∑' n : ℕ, r ^ (n + 1) = r * (1 - r)⁻¹ := by simp only [_root_.pow_succ', ENNReal.tsum_mul_left, ENNReal.tsum_geometric] diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index fe6f17c4e3162..ae6a39293a1df 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -195,8 +195,7 @@ theorem tendsto_div_of_monotone_of_tendsto_div_floor_pow (u : ℕ → ℝ) (l : have H : ∀ n : ℕ, (0 : ℝ) < ⌊c k ^ n⌋₊ := by intro n refine zero_lt_one.trans_le ?_ - simp only [Real.rpow_natCast, Nat.one_le_cast, Nat.one_le_floor_iff, - one_le_pow_of_one_le (cone k).le n] + simp only [Real.rpow_natCast, Nat.one_le_cast, Nat.one_le_floor_iff, one_le_pow₀ (cone k).le] have A : Tendsto (fun n : ℕ => (⌊c k ^ (n + 1)⌋₊ : ℝ) / c k ^ (n + 1) * c k / (⌊c k ^ n⌋₊ / c k ^ n)) atTop (𝓝 (1 * c k / 1)) := by @@ -223,12 +222,12 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by rw [← div_eq_mul_inv, ← div_eq_mul_inv, div_le_div_iff _ (sub_pos.2 hc)] swap - · exact sub_pos.2 (pow_lt_one (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero) + · exact sub_pos.2 (pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero) have : c ^ 3 = c ^ 2 * c := by ring simp only [mul_sub, this, mul_one, inv_pow, sub_le_sub_iff_left] rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel₀ (sq_pos_of_pos cpos).ne', one_mul] simpa using pow_le_pow_right hc.le one_le_two - have C : c⁻¹ ^ 2 < 1 := pow_lt_one (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero + have C : c⁻¹ ^ 2 < 1 := pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero calc (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ ∑ i ∈ Ico ⌊Real.log j / Real.log c⌋₊ N, (1 : ℝ) / (c ^ i) ^ 2 := by @@ -258,7 +257,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc field_simp [(Real.log_pos hc).ne'] ring rw [Real.rpow_sub A, I] - have : c ^ 2 - 1 ≠ 0 := (sub_pos.2 (one_lt_pow hc two_ne_zero)).ne' + have : c ^ 2 - 1 ≠ 0 := (sub_pos.2 (one_lt_pow₀ hc two_ne_zero)).ne' field_simp [hj.ne', (zero_lt_one.trans hc).ne'] ring _ ≤ c ^ 3 * (c - 1)⁻¹ / j ^ 2 := by gcongr @@ -292,7 +291,7 @@ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : rw [mul_div_assoc', mul_one, div_le_div_iff]; rotate_left · apply sq_pos_of_pos refine zero_lt_one.trans_le ?_ - simp only [Nat.le_floor, one_le_pow_of_one_le, hc.le, Nat.one_le_cast, Nat.cast_one] + simp only [Nat.le_floor, one_le_pow₀, hc.le, Nat.one_le_cast, Nat.cast_one] · exact sq_pos_of_pos (pow_pos cpos _) rw [one_mul, ← mul_pow] gcongr diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index bd6d68eeeb2dd..251ede22462dd 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -126,10 +126,10 @@ theorem hundred_le_m [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 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ε) + (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_of_one_le (by norm_num) _ + have h : 1 ≤ 4 ^ P.parts.card := 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] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 4ffe166e5b4b8..1bc7158dd4eef 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -240,7 +240,7 @@ private theorem m_add_one_div_m_le_one_add [Nonempty α] div_eq_mul_one_div _ (49 : ℝ), mul_div_left_comm (2 : ℝ), ← mul_sub_left_distrib, div_pow, div_le_iff₀ (show (0 : ℝ) < ↑100 ^ 2 by norm_num), mul_assoc, sq] refine mul_le_mul_of_nonneg_left ?_ (by sz_positivity) - exact (pow_le_one 5 (by sz_positivity) hε₁).trans (by norm_num) + 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) diff --git a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean index 65daa0a486dab..aace8634bc1a8 100644 --- a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean +++ b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean @@ -201,7 +201,7 @@ lemma eventually_atTop_nonneg_or_nonpos (hf : GrowsPolynomially f) : have le_2n : max n₀ 2 ≤ (2 : ℝ)^n * max n₀ 2 := by nth_rewrite 1 [← one_mul (max n₀ 2)] gcongr - exact one_le_pow_of_one_le (by norm_num : (1 : ℝ) ≤ 2) _ + exact one_le_pow₀ (by norm_num : (1 : ℝ) ≤ 2) have n₀_le_z : n₀ ≤ z := by calc n₀ ≤ max n₀ 2 := by simp _ ≤ (2 : ℝ)^n * max n₀ 2 := le_2n diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index c009cde85a7f2..cd187d52acfae 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1110,7 +1110,7 @@ theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : simp_rw [map_mul, map_pow, map_div₀, abs_natCast] gcongr rw [abv_pow abs] - exact pow_le_one _ (abs.nonneg _) hx + exact pow_le_one₀ (abs.nonneg _) hx _ = abs x ^ n * ∑ m ∈ (range j).filter fun k => n ≤ k, (1 / m.factorial : ℝ) := by simp [abs_mul, abv_pow abs, abs_div, ← mul_sum] _ ≤ abs x ^ n * (n.succ * (n.factorial * n : ℝ)⁻¹) := by @@ -1334,7 +1334,7 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := (calc |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := by gcongr - · exact pow_le_one _ (abs_nonneg _) hx + · exact pow_le_one₀ (abs_nonneg _) hx · rw [sq, ← abs_mul_self, abs_mul] exact mul_le_one hx (abs_nonneg _) hx _ < 1 := by norm_num) diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index 8736c6f964f66..479d737a940f8 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -250,7 +250,7 @@ private lemma I_le (n : ℕ) : I n (π / 2) ≤ 2 := by intros x hx simp only [uIoc_of_le, neg_le_self_iff, zero_le_one, mem_Ioc] at hx rw [norm_eq_abs, abs_mul, abs_pow] - refine mul_le_one (pow_le_one _ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) + refine mul_le_one (pow_le_one₀ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) rw [abs_le] constructor <;> nlinarith diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 94ce55d4cf069..f408bce7cb363 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -452,7 +452,7 @@ theorem exists_orderOf_eq_exponent (hG : ExponentExists G) : ∃ g : G, orderOf rw [(Commute.all _ g).orderOf_mul_eq_mul_orderOf_of_coprime hcoprime, hpk', hg, ha, hk, pow_add, pow_add, pow_one, ← mul_assoc, ← mul_assoc, Nat.div_mul_cancel, mul_assoc, lt_mul_iff_one_lt_right <| hG.orderOf_pos t, ← pow_succ] - · exact one_lt_pow hp.one_lt a.succ_ne_zero + · exact one_lt_pow₀ hp.one_lt a.succ_ne_zero · exact hpk @[to_additive] diff --git a/Mathlib/GroupTheory/FiniteAbelian.lean b/Mathlib/GroupTheory/FiniteAbelian.lean index 0e74393883b9e..8ac8ceb36aec7 100644 --- a/Mathlib/GroupTheory/FiniteAbelian.lean +++ b/Mathlib/GroupTheory/FiniteAbelian.lean @@ -148,7 +148,7 @@ lemma equiv_directSum_zmod_of_finite' (G : Type*) [AddCommGroup G] [Finite G] : refine ⟨{i : ι // n i ≠ 0}, inferInstance, fun i ↦ p i ^ n i, ?_, ⟨e.trans (directSumNeZeroMulEquiv ι _ _).symm⟩⟩ rintro ⟨i, hi⟩ - exact one_lt_pow (hp _).one_lt hi + exact one_lt_pow₀ (hp _).one_lt hi theorem finite_of_fg_torsion [hG' : AddGroup.FG G] (hG : AddMonoid.IsTorsion G) : Finite G := @Module.finite_of_fg_torsion _ _ _ (Module.Finite.iff_addGroup_fg.mpr hG') <| diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 462873bca532e..87db896052260 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -1031,8 +1031,8 @@ theorem orderOf_abs_ne_one (h : |x| ≠ 1) : orderOf x = 0 := by intro n hn hx replace hx : |x| ^ n = 1 := by simpa only [abs_one, abs_pow] using congr_arg abs hx cases' h.lt_or_lt with h h - · exact ((pow_lt_one (abs_nonneg x) h hn.ne').ne hx).elim - · exact ((one_lt_pow h hn.ne').ne' hx).elim + · exact ((pow_lt_one₀ (abs_nonneg x) h hn.ne').ne hx).elim + · exact ((one_lt_pow₀ h hn.ne').ne' hx).elim theorem LinearOrderedRing.orderOf_le_two : orderOf x ≤ 2 := by cases' ne_or_eq |x| 1 with h h diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 48adfd8c9bacd..52fcaf2fb7a60 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -136,7 +136,7 @@ theorem nontrivial_iff_card [Finite G] : Nontrivial G ↔ ∃ n > 0, Nat.card G hk⟩, fun ⟨k, hk0, hk⟩ => Finite.one_lt_card_iff_nontrivial.1 <| - hk.symm ▸ one_lt_pow (Fact.out (p := p.Prime)).one_lt (ne_of_gt hk0)⟩ + hk.symm ▸ one_lt_pow₀ (Fact.out (p := p.Prime)).one_lt (ne_of_gt hk0)⟩ variable {α : Type*} [MulAction G α] diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index eda2a253359f6..38b5ab4ee0b07 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -656,7 +656,7 @@ lemma not_isCyclic_iff_exponent_eq_prime [Group α] {p : ℕ} (hp : p.Prime) let _inst : Fintype α := @Fintype.ofFinite α <| Nat.finite_of_card_ne_zero <| by aesop have hα' : Fintype.card α = p ^ 2 := by simpa using hα have := (Fintype.one_lt_card_iff_nontrivial (α := α)).mp <| - hα' ▸ one_lt_pow hp.one_lt two_ne_zero + hα' ▸ one_lt_pow₀ hp.one_lt two_ne_zero /- in the forward direction, we apply `exponent_eq_prime_iff`, and the reverse direction follows immediately because if `α` has exponent `p`, it has no element of order `p ^ 2`. -/ refine ⟨fun h_cyc ↦ (Monoid.exponent_eq_prime_iff hp).mpr fun g hg ↦ ?_, fun h_exp h_cyc ↦ by diff --git a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean index fb1be755747f1..df01cd8cae6c3 100644 --- a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean +++ b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean @@ -134,7 +134,7 @@ theorem exists_ne_zero_mem_lattice_of_measure_mul_two_pow_le_measure [NormedAddC rw [show μ s < _ ↔ 1 * μ s < _ by rw [one_mul]] refine (mul_lt_mul_right h_mes (ne_of_lt h_cpt.measure_lt_top)).mpr ?_ rw [ofReal_pow (NNReal.coe_nonneg _)] - refine one_lt_pow ?_ (ne_of_gt finrank_pos) + refine one_lt_pow₀ ?_ (ne_of_gt finrank_pos) simp [(exists_seq_strictAnti_tendsto (0 : ℝ≥0)).choose_spec.2.1 n] end MeasureTheory diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 58f251949cb62..39cdf61a38935 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -293,7 +293,7 @@ theorem zeta_sub_one_prime_of_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] Prime (hζ.toInteger - 1) := by letI := IsCyclotomicExtension.numberField {p ^ (k + 1)} ℚ K refine Ideal.prime_of_irreducible_absNorm_span (fun h ↦ ?_) ?_ - · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow hp.out.one_lt (by simp)) + · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow₀ hp.out.one_lt (by simp)) rw [sub_eq_zero] at h simpa using congrArg (algebraMap _ K) h rw [Nat.irreducible_iff_prime, Ideal.absNorm_span_singleton, ← Nat.prime_iff, @@ -312,7 +312,7 @@ theorem zeta_sub_one_prime_of_two_pow [IsCyclotomicExtension {(2 : ℕ+) ^ (k + Prime (hζ.toInteger - 1) := by letI := IsCyclotomicExtension.numberField {(2 : ℕ+) ^ (k + 1)} ℚ K refine Ideal.prime_of_irreducible_absNorm_span (fun h ↦ ?_) ?_ - · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow (by decide) (by simp)) + · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow₀ (by decide) (by simp)) rw [sub_eq_zero] at h simpa using congrArg (algebraMap _ K) h rw [Nat.irreducible_iff_prime, Ideal.absNorm_span_singleton, ← Nat.prime_iff, @@ -453,7 +453,7 @@ theorem not_exists_int_prime_dvd_sub_of_prime_pow_ne_two · simp only [hk, zero_add, pow_one, pow_zero, one_mul, Nat.lt_sub_iff_add_lt, Nat.reduceAdd] at htwo ⊢ exact htwo.symm.lt_of_le hp.1.two_le - · exact one_lt_mul_of_lt_of_le (one_lt_pow hp.1.one_lt hk) + · exact one_lt_mul_of_lt_of_le (one_lt_pow₀ hp.1.one_lt hk) (have := Nat.Prime.two_le hp.out; by omega) rw [sub_eq_iff_eq_add] at h -- We are assuming that `ζ = n + p * x` for some integer `n` and `x : 𝓞 K`. Looking at the @@ -498,7 +498,8 @@ theorem finite_quotient_span_sub_one [hcycl : IsCyclotomicExtension {p ^ (k + 1) have : NumberField K := IsCyclotomicExtension.numberField {p ^ (k + 1)} ℚ K refine Fintype.finite <| Ideal.fintypeQuotientOfFreeOfNeBot _ (fun h ↦ ?_) simp only [Ideal.span_singleton_eq_bot, sub_eq_zero, ← Subtype.coe_inj] at h - exact hζ.ne_one (one_lt_pow hp.1.one_lt (Nat.zero_ne_add_one k).symm) (RingOfIntegers.ext_iff.1 h) + exact hζ.ne_one (one_lt_pow₀ hp.1.one_lt (Nat.zero_ne_add_one k).symm) + (RingOfIntegers.ext_iff.1 h) theorem finite_quotient_span_sub_one' [hcycl : IsCyclotomicExtension {p} ℚ K] (hζ : IsPrimitiveRoot ζ ↑p) : diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index 360f3dc05e0c5..f077508bbee07 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -94,7 +94,7 @@ theorem coprime_of_probablePrime {n b : ℕ} (h : ProbablePrime n b) (h₁ : 1 theorem probablePrime_iff_modEq (n : ℕ) {b : ℕ} (h : 1 ≤ b) : ProbablePrime n b ↔ b ^ (n - 1) ≡ 1 [MOD n] := by - have : 1 ≤ b ^ (n - 1) := one_le_pow_of_one_le h (n - 1) + have : 1 ≤ b ^ (n - 1) := one_le_pow₀ h -- For exact mod_cast rw [Nat.ModEq.comm] constructor diff --git a/Mathlib/NumberTheory/Harmonic/Int.lean b/Mathlib/NumberTheory/Harmonic/Int.lean index cf5d3e429630e..822c1dac331cb 100644 --- a/Mathlib/NumberTheory/Harmonic/Int.lean +++ b/Mathlib/NumberTheory/Harmonic/Int.lean @@ -43,4 +43,4 @@ theorem harmonic_not_int {n : ℕ} (h : 2 ≤ n) : ¬ (harmonic n).isInt := by apply padicNorm.not_int_of_not_padic_int 2 rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (ne_zero_of_lt h)).ne', padicValRat_two_harmonic, neg_neg, zpow_natCast] - exact one_lt_pow one_lt_two (Nat.log_pos one_lt_two h).ne' + exact one_lt_pow₀ one_lt_two (Nat.log_pos one_lt_two h).ne' diff --git a/Mathlib/NumberTheory/Liouville/Basic.lean b/Mathlib/NumberTheory/Liouville/Basic.lean index a9e820e8d2cad..e76ed12616deb 100644 --- a/Mathlib/NumberTheory/Liouville/Basic.lean +++ b/Mathlib/NumberTheory/Liouville/Basic.lean @@ -143,7 +143,7 @@ theorem exists_pos_real_of_irrational_root {α : ℝ} (ha : Irrational α) {f : @exists_one_le_pow_mul_dist ℤ ℕ ℝ _ _ _ (fun y => fR.eval y) α ζ |fR.derivative.eval xm| ?_ z0 (fun y hy => ?_) fun z a hq => ?_ -- 1: the denominators are positive -- essentially by definition; - · exact fun a => one_le_pow_of_one_le ((le_add_iff_nonneg_left 1).mpr a.cast_nonneg) _ + · exact fun a => one_le_pow₀ ((le_add_iff_nonneg_left 1).mpr a.cast_nonneg) -- 2: the polynomial `fR` is Lipschitz at `α` -- as its derivative continuous; · rw [mul_comm] rw [Real.closedBall_eq_Icc] at hy diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean index 2a8dad8fad8f5..3fb42ea6a0f27 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean @@ -183,7 +183,7 @@ theorem liouville_liouvilleNumber {m : ℕ} (hm : 2 ≤ m) : Liouville (liouvill intro n -- the first `n` terms sum to `p / m ^ k!` rcases partialSum_eq_rat (zero_lt_two.trans_le hm) n with ⟨p, hp⟩ - refine ⟨p, m ^ n !, one_lt_pow mZ1 n.factorial_ne_zero, ?_⟩ + refine ⟨p, m ^ n !, one_lt_pow₀ mZ1 n.factorial_ne_zero, ?_⟩ push_cast rw [Nat.cast_pow] at hp -- separate out the sum of the first `n` terms and the rest diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index eed5e96c27d87..fd4fbbd8441a4 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -55,7 +55,7 @@ theorem mersenne_le_mersenne {p q : ℕ} : mersenne p ≤ mersenne q ↔ p ≤ q @[simp] lemma mersenne_odd : ∀ {p : ℕ}, Odd (mersenne p) ↔ p ≠ 0 | 0 => by simp | p + 1 => by - simpa using Nat.Even.sub_odd (one_le_pow_of_one_le one_le_two _) + simpa using Nat.Even.sub_odd (one_le_pow₀ one_le_two) (even_two.pow_of_ne_zero p.succ_ne_zero) odd_one @[simp] theorem mersenne_pos {p : ℕ} : 0 < mersenne p ↔ 0 < p := mersenne_lt_mersenne (p := 0) @@ -87,7 +87,7 @@ theorem one_lt_mersenne {p : ℕ} : 1 < mersenne p ↔ 1 < p := @[simp] theorem succ_mersenne (k : ℕ) : mersenne k + 1 = 2 ^ k := by rw [mersenne, tsub_add_cancel_of_le] - exact one_le_pow_of_one_le (by norm_num) k + exact one_le_pow₀ (by norm_num) namespace LucasLehmer diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index be75dd9a3e776..f5d474fad9d31 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -94,8 +94,7 @@ theorem convexBodyLTFactor_ne_zero : convexBodyLTFactor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) theorem one_le_convexBodyLTFactor : 1 ≤ convexBodyLTFactor K := - one_le_mul (one_le_pow_of_one_le one_le_two _) - (one_le_pow_of_one_le (le_trans one_le_two Real.two_le_pi) _) + one_le_mul (one_le_pow₀ one_le_two) (one_le_pow₀ (one_le_two.trans Real.two_le_pi)) /-- The volume of `(ConvexBodyLt K f)` where `convexBodyLT K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w`. -/ @@ -212,8 +211,7 @@ theorem convexBodyLT'Factor_ne_zero : convexBodyLT'Factor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) theorem one_le_convexBodyLT'Factor : 1 ≤ convexBodyLT'Factor K := - one_le_mul (one_le_pow_of_one_le one_le_two _) - (one_le_pow_of_one_le (le_trans one_le_two Real.two_le_pi) _) + one_le_mul (one_le_pow₀ one_le_two) (one_le_pow₀ (one_le_two.trans Real.two_le_pi)) theorem convexBodyLT'_volume : volume (convexBodyLT' K f w₀) = convexBodyLT'Factor K * ∏ w, (f w) ^ (mult w) := by diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 0800c7995ee2d..b93019fd9a8fb 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -307,7 +307,7 @@ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBd ENNReal.ofReal_one, one_mul, mixedEmbedding.finrank, volume_fundamentalDomain_latticeBasis, 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) + · 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, NNReal.coe_natCast, ← Int.cast_natCast, Int.cast_le] · exact one_le_two diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d2421c79c769e..40c427c9d06f8 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -505,7 +505,7 @@ theorem one_le_of_lt_one {w : InfinitePlace K} {a : (𝓞 K)} (ha : a ≠ 0) rw [← InfinitePlace.prod_eq_abs_norm, ← Finset.prod_const_one] refine Finset.prod_lt_prod_of_nonempty (fun _ _ ↦ ?_) (fun z _ ↦ ?_) Finset.univ_nonempty · exact pow_pos (pos_iff.mpr ((Subalgebra.coe_eq_zero _).not.mpr ha)) _ - · refine pow_lt_one (apply_nonneg _ _) ?_ (by rw [mult]; split_ifs <;> norm_num) + · refine pow_lt_one₀ (apply_nonneg _ _) ?_ (by rw [mult]; split_ifs <;> norm_num) by_cases hz : z = w · rwa [hz] · exact h hz diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index d32fe3c1f17a2..2d3a6893edd99 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -131,7 +131,7 @@ private theorem T_lt_one : T < 1 := by have h := (div_lt_one (deriv_sq_norm_pos hnorm)).2 hnorm rw [T_def]; exact h -private theorem T_pow {n : ℕ} (hn : n ≠ 0) : T ^ n < 1 := pow_lt_one T_nonneg (T_lt_one hnorm) hn +private theorem T_pow {n : ℕ} (hn : n ≠ 0) : T ^ n < 1 := pow_lt_one₀ T_nonneg (T_lt_one hnorm) hn private theorem T_pow' (n : ℕ) : T ^ 2 ^ n < 1 := T_pow hnorm (pow_ne_zero _ two_ne_zero) diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 8c063efdbcb77..8cdd56d82e73b 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -206,7 +206,7 @@ theorem y_ne_zero_of_one_lt_x {a : Solution₁ d} (ha : 1 < a.x) : a.y ≠ 0 := theorem d_pos_of_one_lt_x {a : Solution₁ d} (ha : 1 < a.x) : 0 < d := by refine pos_of_mul_pos_left ?_ (sq_nonneg a.y) rw [a.prop_y, sub_pos] - exact one_lt_pow ha two_ne_zero + exact one_lt_pow₀ ha two_ne_zero /-- If a solution has `x > 1`, then `d` is not a square. -/ theorem d_nonsquare_of_one_lt_x {a : Solution₁ d} (ha : 1 < a.x) : ¬IsSquare d := by diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 4b097fac62731..94bcc41f1f6db 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -457,7 +457,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : refine zero_lt_one.trans_le ?_ apply Nat.le_floor rw [Nat.cast_one] - apply one_le_pow_of_one_le c_one.le + apply one_le_pow₀ c_one.le _ = ENNReal.ofReal (∑ i ∈ range N, Var[S (u i)] / (u i * ε) ^ 2) := by rw [ENNReal.ofReal_sum_of_nonneg fun i _ => ?_] exact div_nonneg (variance_nonneg _ _) (sq_nonneg _) @@ -563,7 +563,7 @@ theorem strong_law_aux6 {c : ℝ} (c_one : 1 < c) : have H : ∀ n : ℕ, (0 : ℝ) < ⌊c ^ n⌋₊ := by intro n refine zero_lt_one.trans_le ?_ - simp only [Nat.one_le_cast, Nat.one_le_floor_iff, one_le_pow_of_one_le c_one.le n] + simp only [Nat.one_le_cast, Nat.one_le_floor_iff, one_le_pow₀ c_one.le] filter_upwards [strong_law_aux4 X hint hindep hident hnonneg c_one, strong_law_aux5 X hint hident hnonneg] with ω hω h'ω rw [← tendsto_sub_nhds_zero_iff, ← Asymptotics.isLittleO_one_iff ℝ] diff --git a/Mathlib/Topology/Algebra/Polynomial.lean b/Mathlib/Topology/Algebra/Polynomial.lean index fb827c4965390..663b8bd81dc20 100644 --- a/Mathlib/Topology/Algebra/Polynomial.lean +++ b/Mathlib/Topology/Algebra/Polynomial.lean @@ -191,8 +191,7 @@ theorem coeff_bdd_of_roots_le {B : ℝ} {d : ℕ} (f : F →+* K) {p : F[X]} (h1 _ ≤ max B 1 ^ d * d.choose (d / 2) := by gcongr; exact (i.choose_mono h3).trans (i.choose_le_middle d) · rw [eq_one_of_roots_le hB h1 h2 h4, Polynomial.map_one, coeff_one] - refine _root_.trans ?_ - (one_le_mul_of_one_le_of_one_le (one_le_pow_of_one_le (le_max_right B 1) d) ?_) + refine le_trans ?_ (one_le_mul_of_one_le_of_one_le (one_le_pow₀ (le_max_right B 1)) ?_) · split_ifs <;> norm_num · exact mod_cast Nat.succ_le_iff.mpr (Nat.choose_pos (d.div_le_self 2)) diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index 80aa1eef70921..f34d6b2a026ac 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -54,7 +54,7 @@ instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by have hVn : ∀ n x, x ∈ Vn n ↔ |Complex.arg x| < Real.pi / 2 ^ (n + 1) := by refine fun n x ↦ ⟨?_, fun hx ↦ ⟨Complex.arg x, hx, Circle.exp_arg x⟩⟩ rintro ⟨t, ht : |t| < _, rfl⟩ - have ht' := ht.trans_le (div_le_self Real.pi_nonneg (one_le_pow_of_one_le one_le_two (n + 1))) + have ht' := ht.trans_le (div_le_self Real.pi_nonneg (one_le_pow₀ one_le_two)) rwa [Circle.arg_exp (neg_lt_of_abs_lt ht') (lt_of_abs_lt ht').le] refine ContinuousMonoidHom.locallyCompactSpace_of_hasBasis Vn ?_ ?_ · intro n x h1 h2 From f6a864f5e846a5eb1dff857031e590d4f7ce0302 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 2 Oct 2024 18:00:54 +0000 Subject: [PATCH 133/174] chore: add a few focusing dots 1 (#17315) More missing cdots found by the multiGoal linter (#12339). --- Mathlib/Algebra/BigOperators/Associated.lean | 16 ++++++++-------- Mathlib/Algebra/Module/Submodule/Bilinear.lean | 16 +++++++++------- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 16 ++++++++-------- .../Order/BigOperators/GroupWithZero/List.lean | 12 ++++++------ Mathlib/Algebra/Order/Floor.lean | 2 +- Mathlib/Analysis/MeanInequalities.lean | 4 ++-- Mathlib/Analysis/PSeries.lean | 4 ++-- .../MorphismProperty/Representable.lean | 4 ++-- 8 files changed, 38 insertions(+), 36 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index ffad576334789..f0cb1a044f328 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -94,14 +94,14 @@ theorem divisor_closure_eq_closure [CancelCommMonoidWithZero α] obtain ⟨ha₁ | ha₂, hs⟩ := hm · rcases ha₁.exists_right_inv with ⟨k, hk⟩ refine hind x (y*k) ?_ hs ?_ - simp only [← mul_assoc, ← hprod, ← Multiset.prod_cons, mul_comm] - refine multiset_prod_mem _ _ (Multiset.forall_mem_cons.2 ⟨subset_closure (Set.mem_def.2 ?_), - Multiset.forall_mem_cons.2 ⟨subset_closure (Set.mem_def.2 ?_), (fun t ht => - subset_closure (hs t ht))⟩⟩) - · left; exact isUnit_of_mul_eq_one_right _ _ hk - · left; exact ha₁ - rw [← mul_one s.prod, ← hk, ← mul_assoc, ← mul_assoc, mul_eq_mul_right_iff, mul_comm] - left; exact hprod + · simp only [← mul_assoc, ← hprod, ← Multiset.prod_cons, mul_comm] + refine multiset_prod_mem _ _ (Multiset.forall_mem_cons.2 ⟨subset_closure (Set.mem_def.2 ?_), + Multiset.forall_mem_cons.2 ⟨subset_closure (Set.mem_def.2 ?_), (fun t ht => + subset_closure (hs t ht))⟩⟩) + · left; exact isUnit_of_mul_eq_one_right _ _ hk + · left; exact ha₁ + · rw [← mul_one s.prod, ← hk, ← mul_assoc, ← mul_assoc, mul_eq_mul_right_iff, mul_comm] + left; exact hprod · rcases ha₂.dvd_mul.1 (Dvd.intro _ hprod) with ⟨c, hc⟩ | ⟨c, hc⟩ · rw [hc]; rw [hc, mul_assoc] at hprod refine Submonoid.mul_mem _ (subset_closure (Set.mem_def.2 ?_)) diff --git a/Mathlib/Algebra/Module/Submodule/Bilinear.lean b/Mathlib/Algebra/Module/Submodule/Bilinear.lean index bacd3bc47d6b5..d6c065576cbbf 100644 --- a/Mathlib/Algebra/Module/Submodule/Bilinear.lean +++ b/Mathlib/Algebra/Module/Submodule/Bilinear.lean @@ -56,13 +56,15 @@ theorem map₂_span_span (f : M →ₗ[R] N →ₗ[R] P) (s : Set M) (t : Set N) apply le_antisymm · rw [map₂_le] apply @span_induction' R M _ _ _ s - intro a ha - apply @span_induction' R N _ _ _ t - intro b hb - exact subset_span ⟨_, ‹_›, _, ‹_›, rfl⟩ - all_goals intros; simp only [*, add_mem, smul_mem, zero_mem, _root_.map_zero, map_add, - LinearMap.zero_apply, LinearMap.add_apply, LinearMap.smul_apply, - map_smul] + on_goal 1 => + intro a ha + apply @span_induction' R N _ _ _ t + · intro b hb + exact subset_span ⟨_, ‹_›, _, ‹_›, rfl⟩ + all_goals + intros + simp only [*, add_mem, smul_mem, zero_mem, _root_.map_zero, map_add, + LinearMap.zero_apply, LinearMap.add_apply, LinearMap.smul_apply, map_smul] · rw [span_le, image2_subset_iff] intro a ha b hb exact apply_mem_map₂ _ (subset_span ha) (subset_span hb) diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index f0273652c337c..a2574ffe76219 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -209,10 +209,10 @@ instance nonUnitalSemiring : NonUnitalSemiring (MonoidAlgebra k G) := mul_assoc := fun f g h => by -- Porting note: `reducible` cannot be `local` so proof gets long. simp only [mul_def] - rw [sum_sum_index]; congr; ext a₁ b₁ - rw [sum_sum_index, sum_sum_index]; congr; ext a₂ b₂ - rw [sum_sum_index, sum_single_index]; congr; ext a₃ b₃ - rw [sum_single_index, mul_assoc, mul_assoc] + rw [sum_sum_index] <;> congr; on_goal 1 => ext a₁ b₁ + rw [sum_sum_index, sum_sum_index] <;> congr; on_goal 1 => ext a₂ b₂ + rw [sum_sum_index, sum_single_index] <;> congr; on_goal 1 => ext a₃ b₃ + on_goal 1 => rw [sum_single_index, mul_assoc, mul_assoc] all_goals simp only [single_zero, single_add, forall_true_iff, add_mul, mul_add, zero_mul, mul_zero, sum_zero, sum_add] } @@ -974,10 +974,10 @@ instance nonUnitalSemiring : NonUnitalSemiring k[G] := mul_assoc := fun f g h => by -- Porting note: `reducible` cannot be `local` so proof gets long. simp only [mul_def] - rw [sum_sum_index]; congr; ext a₁ b₁ - rw [sum_sum_index, sum_sum_index]; congr; ext a₂ b₂ - rw [sum_sum_index, sum_single_index]; congr; ext a₃ b₃ - rw [sum_single_index, mul_assoc, add_assoc] + rw [sum_sum_index] <;> congr; on_goal 1 => ext a₁ b₁ + rw [sum_sum_index, sum_sum_index] <;> congr; on_goal 1 => ext a₂ b₂ + rw [sum_sum_index, sum_single_index] <;> congr; on_goal 1 => ext a₃ b₃ + on_goal 1 => rw [sum_single_index, mul_assoc, add_assoc] all_goals simp only [single_zero, single_add, forall_true_iff, add_mul, mul_add, zero_mul, mul_zero, sum_zero, sum_add] } diff --git a/Mathlib/Algebra/Order/BigOperators/GroupWithZero/List.lean b/Mathlib/Algebra/Order/BigOperators/GroupWithZero/List.lean index 77f4d768aee3b..e38120c46d7ac 100644 --- a/Mathlib/Algebra/Order/BigOperators/GroupWithZero/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/GroupWithZero/List.lean @@ -51,12 +51,12 @@ theorem prod_map_le_prod_map₀ {ι : Type*} {s : List ι} (f : ι → R) (g : · intro i hi apply h simp [hi] - apply prod_nonneg - · simp only [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + · apply prod_nonneg + simp only [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] intro a ha apply h0 simp [ha] - apply (h0 _ _).trans (h _ _) <;> simp + · apply (h0 _ _).trans (h _ _) <;> simp omit [PosMulMono R] variable [PosMulStrictMono R] [NeZero (1 : R)] @@ -89,11 +89,11 @@ theorem prod_map_lt_prod_map {ι : Type*} {s : List ι} (hs : s ≠ []) apply le_of_lt apply h simp [hi] - apply prod_pos - · simp only [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + · apply prod_pos + simp only [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] intro a ha apply h0 simp [ha] - apply le_of_lt ((h0 _ _).trans (h _ _)) <;> simp + · apply le_of_lt ((h0 _ _).trans (h _ _)) <;> simp end List diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 7d9f071cca4ea..a3689333018fb 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -526,7 +526,7 @@ lemma ceil_lt_mul (hb : 1 < b) (hba : ⌈(b - 1)⁻¹⌉₊ / b < a) : ⌈a⌉ lemma ceil_le_mul (hb : 1 < b) (hba : ⌈(b - 1)⁻¹⌉₊ / b ≤ a) : ⌈a⌉₊ ≤ b * a := by obtain rfl | hba := hba.eq_or_lt · rw [mul_div_cancel₀, cast_le, ceil_le] - exact _root_.div_le_self (by positivity) hb.le + · exact _root_.div_le_self (by positivity) hb.le · positivity · exact (ceil_lt_mul hb hba).le diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 6f6a2a5350282..d8e1833217bd0 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -668,8 +668,8 @@ lemma compact_inner_le_weight_mul_Lp_of_nonneg (s : Finset ι) {p : ℝ} (hp : 1 𝔼 i ∈ s, w i * f i ≤ (𝔼 i ∈ s, w i) ^ (1 - p⁻¹) * (𝔼 i ∈ s, w i * f i ^ p) ^ p⁻¹ := by simp_rw [expect_eq_sum_div_card] rw [div_rpow, div_rpow, div_mul_div_comm, ← rpow_add', sub_add_cancel, rpow_one] - gcongr - · exact inner_le_weight_mul_Lp_of_nonneg s hp _ _ hw hf + · gcongr + exact inner_le_weight_mul_Lp_of_nonneg s hp _ _ hw hf any_goals simp · exact sum_nonneg fun i _ ↦ by have := hw i; have := hf i; positivity · exact sum_nonneg fun i _ ↦ by have := hw i; positivity diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 29941c56c0472..aed28d9ca5915 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -162,8 +162,8 @@ theorem tsum_schlomilch_le {C : ℕ} (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → le_trans ?_ (add_le_add_left (mul_le_mul_of_nonneg_left (ENNReal.sum_le_tsum <| Finset.Ico (u 0 + 1) (u n + 1)) ?_) _) - simpa using Finset.sum_schlomilch_le hf h_pos h_nonneg hu h_succ_diff n - exact zero_le _ + · simpa using Finset.sum_schlomilch_le hf h_pos h_nonneg hu h_succ_diff n + · exact zero_le _ theorem tsum_condensed_le (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) : (∑' k : ℕ, 2 ^ k * f (2 ^ k)) ≤ f 1 + 2 * ∑' k, f k := by diff --git a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean index ba9ba8a06a053..91bf09c44907e 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Representable.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Representable.lean @@ -280,8 +280,8 @@ instance isMultiplicative : IsMultiplicative F.relativelyRepresentable where lemma stableUnderBaseChange : StableUnderBaseChange F.relativelyRepresentable := by intro X Y Y' X' f g f' g' P₁ hg a h refine ⟨hg.pullback (h ≫ f), hg.snd (h ≫ f), ?_, ?_⟩ - apply P₁.lift (hg.fst (h ≫ f)) (F.map (hg.snd (h ≫ f)) ≫ h) (by simpa using hg.w (h ≫ f)) - apply IsPullback.of_right' (hg.isPullback (h ≫ f)) P₁ + · apply P₁.lift (hg.fst (h ≫ f)) (F.map (hg.snd (h ≫ f)) ≫ h) (by simpa using hg.w (h ≫ f)) + · apply IsPullback.of_right' (hg.isPullback (h ≫ f)) P₁ instance respectsIso : RespectsIso F.relativelyRepresentable := (stableUnderBaseChange F).respectsIso From bf960d14a779837ce19ad56aa5d472b300e133cd Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 2 Oct 2024 18:00:55 +0000 Subject: [PATCH 134/174] chore(Order/Interval/Set/Basic): remove porting note (#17321) Co-authored-by: Moritz Firsching --- Mathlib/Order/Interval/Set/Basic.lean | 15 +++++---------- 1 file changed, 5 insertions(+), 10 deletions(-) diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 26d1511237b73..07d4cb76870cc 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -1566,16 +1566,11 @@ theorem Ioc_union_Ioc_symm : Ioc a b ∪ Ioc b a = Ioc (min a b) (max a b) := by @[simp] theorem Ioc_union_Ioc_union_Ioc_cycle : Ioc a b ∪ Ioc b c ∪ Ioc c a = Ioc (min a (min b c)) (max a (max b c)) := by - rw [Ioc_union_Ioc, Ioc_union_Ioc] <;> - -- Porting note: mathlib3 proof finished from here as follows: - -- (It can probably be restored after https://github.com/leanprover-community/mathlib4/pull/856) - -- ac_rfl - -- all_goals - -- solve_by_elim (config := { max_depth := 5 }) [min_le_of_left_le, min_le_of_right_le, - -- le_max_of_le_left, le_max_of_le_right, le_refl] - simp [min_le_of_left_le, min_le_of_right_le, le_max_of_le_left, le_max_of_le_right, le_refl, - min_assoc, max_comm] - + rw [Ioc_union_Ioc, Ioc_union_Ioc] + ac_rfl + all_goals + solve_by_elim (config := { maxDepth := 5 }) [min_le_of_left_le, min_le_of_right_le, + le_max_of_le_left, le_max_of_le_right, le_refl] end LinearOrder /-! From 551dfed37fb761cc7bc33eac160eae1771191226 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 2 Oct 2024 18:00:56 +0000 Subject: [PATCH 135/174] chore: add a focusing dot (#17346) --- Mathlib/RingTheory/LaurentSeries.lean | 2 +- Mathlib/RingTheory/TwoSidedIdeal/Operations.lean | 3 ++- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index 669be5c1d5cd1..eea4d59465c7d 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -771,7 +771,7 @@ theorem Cauchy.eventually_mem_nhds {ℱ : Filter (LaurentSeries K)} (hℱ : Cauc rw [← WithZero.coe_unzero γ.ne_zero, WithZero.coe_lt_coe, hD₀, neg_neg, ofAdd_sub, ofAdd_toAdd, div_lt_comm, div_self', ← ofAdd_zero, Multiplicative.ofAdd_lt] exact zero_lt_one - apply coeff_eventually_equal hℱ |>.mono + apply coeff_eventually_equal (D := D) hℱ |>.mono intro _ hf apply lt_of_le_of_lt (valuation_le_iff_coeff_lt_eq_zero K |>.mpr _) hD intro n hn diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean index ea08a0c3e8866..455172552a77c 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean @@ -106,7 +106,8 @@ def ker : TwoSidedIdeal R := (by rintro _ _ (h : f _ = 0); simp [h]) lemma mem_ker {x : R} : x ∈ ker f ↔ f x = 0 := by - delta ker; rw [mem_mk']; rfl + delta ker; rw [mem_mk'] + · rfl · rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2] · rintro _ (h : f _ = 0); simp [h] · rintro _ _ (h : f _ = 0); simp [h] From bb7a404341c1b8f113a011a0a60b367b7d2649ff Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 2 Oct 2024 19:31:14 +0000 Subject: [PATCH 136/174] feat(AlgebraicGeometry/Tilde): add the map from `M` localising at `x` to the stalk of `M^~` at x (#14809) 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: Weihong Xu Co-authored-by: morel --- Mathlib/Algebra/Module/LocalizedModule.lean | 23 +++++ Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 100 ++++++++++++++++++- 2 files changed, 122 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index 45ebbceedbb8a..5edf2a007db89 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -402,6 +402,29 @@ noncomputable instance isModule' : Module R (LocalizedModule S M) := theorem smul'_mk (r : R) (s : S) (m : M) : r • mk m s = mk (r • m) s := by erw [mk_smul_mk r m 1 s, one_mul] +lemma smul_eq_iff_of_mem + (r : R) (hr : r ∈ S) (x y : LocalizedModule S M) : + r • x = y ↔ x = Localization.mk 1 ⟨r, hr⟩ • y := by + induction x using induction_on with + | h m s => + induction y using induction_on with + | h n t => + rw [smul'_mk, mk_smul_mk, one_smul, mk_eq, mk_eq] + simp only [Subtype.exists, Submonoid.mk_smul, exists_prop] + fconstructor + · rintro ⟨a, ha, eq1⟩ + refine ⟨a, ha, ?_⟩ + rw [mul_smul, ← eq1, Submonoid.mk_smul, smul_comm r t] + · rintro ⟨a, ha, eq1⟩ + refine ⟨a, ha, ?_⟩ + rw [← eq1, mul_comm, mul_smul, Submonoid.mk_smul] + rfl + +lemma eq_zero_of_smul_eq_zero + (r : R) (hr : r ∈ S) (x : LocalizedModule S M) (hx : r • x = 0) : x = 0 := by + rw [smul_eq_iff_of_mem (hr := hr)] at hx + rw [hx, smul_zero] + theorem smul'_mul {A : Type*} [Semiring A] [Algebra R A] (x : T) (p₁ p₂ : LocalizedModule S A) : x • p₁ * p₂ = x • (p₁ * p₂) := by induction p₁, p₂ using induction_on₂ with | _ a₁ s₁ a₂ s₂ => _ diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index b52b243e1ff3c..b81f47d19763c 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -1,13 +1,15 @@ /- Copyright (c) 2024 Weihong Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Weihong Xu +Authors: Kevin Buzzard, Johan Commelin, Amelia Livingston, Sophie Morel, Jujian Zhang, Weihong Xu -/ import Mathlib.Algebra.Module.LocalizedModule import Mathlib.AlgebraicGeometry.StructureSheaf import Mathlib.AlgebraicGeometry.Modules.Sheaf import Mathlib.Algebra.Category.ModuleCat.Sheaf +import Mathlib.Algebra.Category.ModuleCat.FilteredColimits +import Mathlib.CategoryTheory.Limits.ConcreteCategory.WithAlgebraicStructures /-! @@ -21,6 +23,12 @@ 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. +## 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`. + -/ universe u @@ -155,4 +163,94 @@ noncomputable def tilde : (Spec (CommRingCat.of R)).Modules where isSheaf := (TopCat.Presheaf.isSheaf_iff_isSheaf_comp (forget AddCommGrp) _ ).2 M.tildeInType.2 +/-- +This is `M^~` as a sheaf of `R`-modules. +-/ +noncomputable def tildeInModuleCat : + TopCat.Presheaf (ModuleCat R) (PrimeSpectrum.Top R) := + (PresheafOfModules.forgetToPresheafModuleCat (op ⊤) <| + Limits.initialOpOfTerminal Limits.isTerminalTop).obj (tilde M).1 ⋙ + ModuleCat.restrictScalars (StructureSheaf.globalSectionsIso R).hom + +namespace Tilde + +@[simp] +theorem res_apply (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) + (s : (tildeInModuleCat M).obj (op U)) (x : V) : + ((tildeInModuleCat M).map i.op s).1 x = (s.1 (i x) : _) := + rfl + +lemma smul_section_apply (r : R) (U : Opens (PrimeSpectrum.Top R)) + (s : (tildeInModuleCat M).1.obj (op U)) (x : U) : + (r • s).1 x = r • (s.1 x) := rfl + +lemma smul_stalk_no_nonzero_divisor {x : PrimeSpectrum R} + (r : x.asIdeal.primeCompl) (st : (tildeInModuleCat M).stalk x) (hst : r.1 • st = 0) : + st = 0 := by + refine Limits.Concrete.colimit_no_zero_smul_divisor (hx := hst) + ⟨op ⟨PrimeSpectrum.basicOpen r.1, r.2⟩, fun U i s hs ↦ Subtype.eq <| funext fun pt ↦ ?_⟩ + exact LocalizedModule.eq_zero_of_smul_eq_zero (hx := congr_fun (Subtype.ext_iff.1 hs) pt) <| + i.unop pt |>.2 + +/-- +If `U` is an open subset of `Spec R`, this is the morphism of `R`-modules from `M` to +`M^~(U)`. +-/ +def toOpen (U : Opens (PrimeSpectrum.Top R)) : + ModuleCat.of R M ⟶ (tildeInModuleCat M).1.obj (op U) where + toFun f := + ⟨fun x ↦ LocalizedModule.mkLinearMap _ _ f, fun x ↦ + ⟨U, x.2, 𝟙 _, f, 1, fun y ↦ ⟨(Ideal.ne_top_iff_one _).1 y.1.2.1, by simp⟩⟩⟩ + map_add' f g := Subtype.eq <| funext fun x ↦ LinearMap.map_add _ _ _ + map_smul' r m := by + simp only [isLocallyFraction_pred, LocalizedModule.mkLinearMap_apply, LinearMapClass.map_smul, + RingHom.id_apply] + rfl + +@[simp] +theorem toOpen_res (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) : + toOpen M U ≫ (tildeInModuleCat M).map i.op = toOpen M V := + rfl + +/-- +If `x` is a point of `Spec R`, this is the morphism of `R`-modules from `M` to the stalk of +`M^~` at `x`. +-/ +noncomputable def toStalk (x : PrimeSpectrum.Top R) : + ModuleCat.of R M ⟶ TopCat.Presheaf.stalk (tildeInModuleCat M) x := + (toOpen M ⊤ ≫ TopCat.Presheaf.germ (tildeInModuleCat M) ⟨x, by trivial⟩) + +open LocalizedModule TopCat.Presheaf in +lemma isUnit_toStalk (x : PrimeSpectrum.Top R) (r : x.asIdeal.primeCompl) : + IsUnit ((algebraMap R (Module.End R ((tildeInModuleCat M).stalk x))) r) := by + rw [Module.End_isUnit_iff] + refine ⟨LinearMap.ker_eq_bot.1 <| eq_bot_iff.2 fun st (h : r.1 • st = 0) ↦ + smul_stalk_no_nonzero_divisor M r st h, fun st ↦ ?_⟩ + obtain ⟨U, mem, s, rfl⟩ := germ_exist (F := M.tildeInModuleCat) x st + let O := U ⊓ (PrimeSpectrum.basicOpen r) + refine ⟨germ M.tildeInModuleCat (⟨x, ⟨mem, r.2⟩⟩ : O) + ⟨fun q ↦ (Localization.mk 1 ⟨r, q.2.2⟩ : Localization.AtPrime q.1.asIdeal) • s.1 + ⟨q.1, q.2.1⟩, fun q ↦ ?_⟩, by + simpa only [Module.algebraMap_end_apply, ← map_smul] using + germ_ext (W := O) (hxW := ⟨mem, r.2⟩) (iWU := 𝟙 _) (iWV := homOfLE inf_le_left) _ <| + Subtype.eq <| funext fun y ↦ smul_eq_iff_of_mem (S := y.1.1.primeCompl) r _ _ _ |>.2 rfl⟩ + obtain ⟨V, mem_V, iV, num, den, hV⟩ := s.2 ⟨q.1, q.2.1⟩ + refine ⟨V ⊓ O, ⟨mem_V, q.2⟩, homOfLE inf_le_right, num, r * den, fun y ↦ ?_⟩ + obtain ⟨h1, h2⟩ := hV ⟨y, y.2.1⟩ + refine ⟨y.1.asIdeal.primeCompl.mul_mem y.2.2.2 h1, ?_⟩ + simp only [Opens.coe_inf, isLocallyFraction_pred, mkLinearMap_apply, + smul_eq_iff_of_mem (S := y.1.1.primeCompl) (hr := h1), mk_smul_mk, one_smul, mul_one] at h2 ⊢ + simpa only [h2, mk_smul_mk, one_smul, smul'_mk, mk_eq] using ⟨1, by simp only [one_smul]; rfl⟩ + +/-- +The morphism of `R`-modules from the localization of `M` at the prime ideal corresponding to `x` +to the stalk of `M^~` at `x`. +-/ +noncomputable def localizationToStalk (x : PrimeSpectrum.Top R) : + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) ⟶ + (TopCat.Presheaf.stalk (tildeInModuleCat M) x) := + LocalizedModule.lift _ (toStalk M x) <| isUnit_toStalk M x + +end Tilde + end ModuleCat From 3d1891294a676adce1919104ebc44b82fdb2a24e Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 2 Oct 2024 20:09:38 +0000 Subject: [PATCH 137/174] chore: remove `.val` where elaboration works (#17363) In all but two places where #6607 is mentioned the elaboration now works correctly. Not sure why it does not for `det_units_conj` and `det_units_conj'`, hence keeping the TODO, but changing to `.val` there because it is a little less verbose. Co-authored-by: Moritz Firsching --- Mathlib/Algebra/Lie/SkewAdjoint.lean | 3 +-- Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean | 6 ++---- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 8 ++++---- Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean | 3 +-- 4 files changed, 8 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/Lie/SkewAdjoint.lean b/Mathlib/Algebra/Lie/SkewAdjoint.lean index bf72da94b1715..4ba8facc11fda 100644 --- a/Mathlib/Algebra/Lie/SkewAdjoint.lean +++ b/Mathlib/Algebra/Lie/SkewAdjoint.lean @@ -126,10 +126,9 @@ def skewAdjointMatricesLieSubalgebraEquiv (P : Matrix n n R) (h : Invertible P) exact this simp [Matrix.IsSkewAdjoint, J.isAdjointPair_equiv _ _ P (isUnit_of_invertible P)] --- TODO(mathlib4#6607): fix elaboration so annotation on `A` isn't needed theorem skewAdjointMatricesLieSubalgebraEquiv_apply (P : Matrix n n R) (h : Invertible P) (A : skewAdjointMatricesLieSubalgebra J) : - ↑(skewAdjointMatricesLieSubalgebraEquiv J P h A) = P⁻¹ * (A : Matrix n n R) * P := by + ↑(skewAdjointMatricesLieSubalgebraEquiv J P h A) = P⁻¹ * A * P := by simp [skewAdjointMatricesLieSubalgebraEquiv] /-- An equivalence of matrix algebras commuting with the transpose endomorphisms restricts to an diff --git a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean index 2f0463ab0aa36..a481e6f7c4823 100644 --- a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean @@ -145,17 +145,15 @@ nonrec theorem isUnit_exp (A : Matrix m m 𝔸) : IsUnit (exp 𝕂 A) := by letI : NormedAlgebra 𝕂 (Matrix m m 𝔸) := Matrix.linftyOpNormedAlgebra exact isUnit_exp _ A --- TODO(mathlib4#6607): fix elaboration so `val` isn't needed nonrec theorem exp_units_conj (U : (Matrix m m 𝔸)ˣ) (A : Matrix m m 𝔸) : - exp 𝕂 (U.val * A * (U⁻¹).val) = U.val * exp 𝕂 A * (U⁻¹).val := by + exp 𝕂 (U * A * U⁻¹) = U * exp 𝕂 A * U⁻¹ := by letI : SeminormedRing (Matrix m m 𝔸) := Matrix.linftyOpSemiNormedRing letI : NormedRing (Matrix m m 𝔸) := Matrix.linftyOpNormedRing letI : NormedAlgebra 𝕂 (Matrix m m 𝔸) := Matrix.linftyOpNormedAlgebra exact exp_units_conj _ U A --- TODO(mathlib4#6607): fix elaboration so `val` isn't needed theorem exp_units_conj' (U : (Matrix m m 𝔸)ˣ) (A : Matrix m m 𝔸) : - exp 𝕂 ((U⁻¹).val * A * U.val) = (U⁻¹).val * exp 𝕂 A * U.val := + exp 𝕂 (U⁻¹ * A * U) = U⁻¹ * exp 𝕂 A * U := exp_units_conj 𝕂 U⁻¹ A end Normed diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index f29c7888f644b..9490aa658bb95 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -183,14 +183,14 @@ theorem det_mul_left_comm (M N P : Matrix m m R) : det (M * (N * P)) = det (N * theorem det_mul_right_comm (M N P : Matrix m m R) : det (M * N * P) = det (M * P * N) := by rw [Matrix.mul_assoc, Matrix.mul_assoc, det_mul, det_mul_comm N P, ← det_mul] --- TODO(mathlib4#6607): fix elaboration so that the ascription isn't needed +-- TODO(mathlib4#6607): fix elaboration so `val` isn't needed theorem det_units_conj (M : (Matrix m m R)ˣ) (N : Matrix m m R) : - det ((M : Matrix _ _ _) * N * (↑M⁻¹ : Matrix _ _ _)) = det N := by + det (M.val * N * M⁻¹.val) = det N := by rw [det_mul_right_comm, Units.mul_inv, one_mul] --- TODO(mathlib4#6607): fix elaboration so that the ascription isn't needed +-- TODO(mathlib4#6607): fix elaboration so `val` isn't needed theorem det_units_conj' (M : (Matrix m m R)ˣ) (N : Matrix m m R) : - det ((↑M⁻¹ : Matrix _ _ _) * N * (↑M : Matrix _ _ _)) = det N := + det (M⁻¹.val * N * ↑M.val) = det N := det_units_conj M⁻¹ N /-- Transposing a matrix preserves the determinant. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index 13e2bba24b39b..8e620f838f3f0 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -559,8 +559,7 @@ theorem Matrix.isAdjointPair_equiv (P : Matrix n n R) (h : IsUnit P) : let v := Pᵀ.nonsingInvUnit (P.isUnit_det_transpose h') let x := A₁ᵀ * Pᵀ * J let y := J * P * A₂ - -- TODO(mathlib4#6607): fix elaboration so `val` isn't needed - suffices x * u.val = v.val * y ↔ (v⁻¹).val * x = y * (u⁻¹).val by + suffices x * u = v * y ↔ v⁻¹ * x = y * u⁻¹ by dsimp only [Matrix.IsAdjointPair] simp only [Matrix.transpose_mul] simp only [← mul_assoc, P.transpose_nonsing_inv] From 3c0774cd67ade0df6ba78798c0edd7b1f9e4fd76 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 20:35:32 +0000 Subject: [PATCH 138/174] refactor(ENNReal): detopologise sSup/sInf lemmas (#16742) These lemmas can be proved much earlier with no topology. --- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 4 +- Mathlib/Analysis/Oscillation.lean | 3 +- Mathlib/Data/ENNReal/Inv.lean | 228 ++++++++++++++++++ Mathlib/Data/ENNReal/Real.lean | 52 +--- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 3 +- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 2 +- .../OuterMeasure/OfFunction.lean | 2 +- Mathlib/Topology/EMetricSpace/Diam.lean | 1 - Mathlib/Topology/Instances/ENNReal.lean | 147 +---------- 9 files changed, 258 insertions(+), 184 deletions(-) diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index f967929332fbd..61bab9e550cc3 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -203,11 +203,9 @@ theorem spectralRadius_le_liminf_pow_nnnorm_pow_one_div (a : A) : refine ENNReal.le_of_forall_lt_one_mul_le fun ε hε => ?_ by_cases h : ε = 0 · simp only [h, zero_mul, zero_le'] - have hε' : ε⁻¹ ≠ ∞ := fun h' => - h (by simpa only [inv_inv, inv_top] using congr_arg (fun x : ℝ≥0∞ => x⁻¹) h') simp only [ENNReal.mul_le_iff_le_inv h (hε.trans_le le_top).ne, mul_comm ε⁻¹, liminf_eq_iSup_iInf_of_nat', ENNReal.iSup_mul] - conv_rhs => arg 1; intro i; rw [ENNReal.iInf_mul hε'] + conv_rhs => arg 1; intro i; rw [ENNReal.iInf_mul (by simp [h])] rw [← ENNReal.inv_lt_inv, inv_one] at hε obtain ⟨N, hN⟩ := eventually_atTop.mp (ENNReal.eventually_pow_one_div_le (ENNReal.coe_ne_top : ↑‖(1 : A)‖₊ ≠ ∞) hε) diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index 16bc8089b08b1..c6327e3b9a059 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 James Sundstrom. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: James Sundstrom -/ -import Mathlib.Topology.EMetricSpace.Diam +import Mathlib.Data.ENNReal.Real import Mathlib.Order.WellFoundedSet +import Mathlib.Topology.EMetricSpace.Diam /-! # Oscillation diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 1be528c876575..5847cba382776 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -601,5 +601,233 @@ protected theorem zpow_sub {x : ℝ≥0∞} (x_ne_zero : x ≠ 0) (x_ne_top : x x ^ (m - n) = (x ^ m) * (x ^ n)⁻¹ := by rw [sub_eq_add_neg, ENNReal.zpow_add x_ne_zero x_ne_top, ENNReal.zpow_neg x_ne_zero x_ne_top n] +variable {ι κ : Sort*} {f g : ι → ℝ≥0∞} {s : Set ℝ≥0∞} {a : ℝ≥0∞} + +@[simp] lemma iSup_eq_zero : ⨆ i, f i = 0 ↔ ∀ i, f i = 0 := iSup_eq_bot + +@[simp] lemma iSup_zero_eq_zero : ⨆ _ : ι, (0 : ℝ≥0∞) = 0 := by simp + +lemma iSup_natCast : ⨆ n : ℕ, (n : ℝ≥0∞) = ∞ := + (iSup_eq_top _).2 fun _b hb => ENNReal.exists_nat_gt (lt_top_iff_ne_top.1 hb) + +@[simp] lemma iSup_lt_eq_self (a : ℝ≥0∞) : ⨆ b, ⨆ _ : b < a, b = a := by + refine le_antisymm (iSup₂_le fun b hb ↦ hb.le) ?_ + refine le_of_forall_lt fun c hca ↦ ?_ + obtain ⟨d, hcd, hdb⟩ := exists_between hca + exact hcd.trans_le <| le_iSup₂_of_le d hdb le_rfl + +lemma isUnit_iff : IsUnit a ↔ a ≠ 0 ∧ a ≠ ∞ := by + refine ⟨fun ha ↦ ⟨ha.ne_zero, ?_⟩, + fun ha ↦ ⟨⟨a, a⁻¹, ENNReal.mul_inv_cancel ha.1 ha.2, ENNReal.inv_mul_cancel ha.1 ha.2⟩, rfl⟩⟩ + obtain ⟨u, rfl⟩ := ha + rintro hu + have := congr($hu * u⁻¹) + norm_cast at this + simp [mul_inv_cancel] at this + +/-- Left multiplication by a nonzero finite `a` as an order isomorphism. -/ +@[simps! toEquiv apply symm_apply] +def mulLeftOrderIso (a : ℝ≥0∞) (ha : IsUnit a) : ℝ≥0∞ ≃o ℝ≥0∞ where + toEquiv := ha.unit.mulLeft + map_rel_iff' := by simp [ENNReal.mul_le_mul_left, ha.ne_zero, (isUnit_iff.1 ha).2] + +/-- Right multiplication by a nonzero finite `a` as an order isomorphism. -/ +@[simps! toEquiv apply symm_apply] +def mulRightOrderIso (a : ℝ≥0∞) (ha : IsUnit a) : ℝ≥0∞ ≃o ℝ≥0∞ where + toEquiv := ha.unit.mulRight + map_rel_iff' := by simp [ENNReal.mul_le_mul_right, ha.ne_zero, (isUnit_iff.1 ha).2] + +lemma mul_iSup (a : ℝ≥0∞) (f : ι → ℝ≥0∞) : a * ⨆ i, f i = ⨆ i, a * f i := by + by_cases hf : ∀ i, f i = 0 + · simp [hf] + obtain rfl | ha₀ := eq_or_ne a 0 + · simp + obtain rfl | ha := eq_or_ne a ∞ + · obtain ⟨i, hi⟩ := not_forall.1 hf + simpa [iSup_eq_zero.not.2 hf, eq_comm (a := ⊤)] using le_iSup_of_le i (top_mul hi).ge + · exact (mulLeftOrderIso _ <| isUnit_iff.2 ⟨ha₀, ha⟩).map_iSup _ + +lemma iSup_mul (f : ι → ℝ≥0∞) (a : ℝ≥0∞) : (⨆ i, f i) * a = ⨆ i, f i * a := by + simp [mul_comm, mul_iSup] + +lemma mul_sSup {a : ℝ≥0∞} : a * sSup s = ⨆ b ∈ s, a * b := by + simp only [sSup_eq_iSup, mul_iSup] + +lemma sSup_mul {a : ℝ≥0∞} : sSup s * a = ⨆ b ∈ s, b * a := by + simp only [sSup_eq_iSup, iSup_mul] + +lemma iSup_div (f : ι → ℝ≥0∞) (a : ℝ≥0∞) : iSup f / a = ⨆ i, f i / a := iSup_mul .. +lemma sSup_div (s : Set ℝ≥0∞) (a : ℝ≥0∞) : sSup s / a = ⨆ b ∈ s, b / a := sSup_mul .. + +/-- Very general version for distributivity of multiplication over an infimum. + +See `ENNReal.mul_iInf_of_ne` for the special case assuming `a ≠ 0` and `a ≠ ∞`, and +`ENNReal.mul_iInf` for the special case assuming `Nonempty ι`. -/ +lemma mul_iInf' (hinfty : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) (h₀ : a = 0 → Nonempty ι) : + a * ⨅ i, f i = ⨅ i, a * f i := by + obtain rfl | ha₀ := eq_or_ne a 0 + · simp [h₀ rfl] + obtain rfl | ha := eq_or_ne a ∞ + · obtain ⟨i, hi⟩ | hf := em (∃ i, f i = 0) + · rw [(iInf_eq_bot _).2, (iInf_eq_bot _).2, bot_eq_zero, mul_zero] <;> + exact fun _ _↦ ⟨i, by simpa [hi]⟩ + · rw [top_mul (mt (hinfty rfl) hf), eq_comm, iInf_eq_top] + exact fun i ↦ top_mul fun hi ↦ hf ⟨i, hi⟩ + · exact (mulLeftOrderIso _ <| isUnit_iff.2 ⟨ha₀, ha⟩).map_iInf _ + +/-- Very general version for distributivity of multiplication over an infimum. + +See `ENNReal.iInf_mul_of_ne` for the special case assuming `a ≠ 0` and `a ≠ ∞`, and +`ENNReal.iInf_mul` for the special case assuming `Nonempty ι`. -/ +lemma iInf_mul' (hinfty : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) (h₀ : a = 0 → Nonempty ι) : + (⨅ i, f i) * a = ⨅ i, f i * a := by simpa only [mul_comm a] using mul_iInf' hinfty h₀ + +/-- If `a ≠ 0` and `a ≠ ∞`, then right multiplication by `a` maps infimum to infimum. + +See `ENNReal.mul_iInf'` for the general case, and `ENNReal.iInf_mul` for another special case that +assumes `Nonempty ι` but does not require `a ≠ 0`, and `ENNReal`. -/ +lemma mul_iInf_of_ne (ha₀ : a ≠ 0) (ha : a ≠ ∞) : a * ⨅ i, f i = ⨅ i, a * f i := + mul_iInf' (by simp [ha]) (by simp [ha₀]) + +/-- If `a ≠ 0` and `a ≠ ∞`, then right multiplication by `a` maps infimum to infimum. + +See `ENNReal.iInf_mul'` for the general case, and `ENNReal.iInf_mul` for another special case that +assumes `Nonempty ι` but does not require `a ≠ 0`. -/ +lemma iInf_mul_of_ne (ha₀ : a ≠ 0) (ha : a ≠ ∞) : (⨅ i, f i) * a = ⨅ i, f i * a := + iInf_mul' (by simp [ha]) (by simp [ha₀]) + +/-- See `ENNReal.mul_iInf'` for the general case, and `ENNReal.mul_iInf_of_ne` for another special +case that assumes `a ≠ 0` but does not require `Nonempty ι`. -/ +lemma mul_iInf [Nonempty ι] (hinfty : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) : + a * ⨅ i, f i = ⨅ i, a * f i := mul_iInf' hinfty fun _ ↦ ‹Nonempty ι› + +/-- See `ENNReal.iInf_mul'` for the general case, and `ENNReal.iInf_mul_of_ne` for another special +case that assumes `a ≠ 0` but does not require `Nonempty ι`. -/ +lemma iInf_mul [Nonempty ι] (hinfty : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) : + (⨅ i, f i) * a = ⨅ i, f i * a := iInf_mul' hinfty fun _ ↦ ‹Nonempty ι› + +/-- Very general version for distributivity of division over an infimum. + +See `ENNReal.iInf_div_of_ne` for the special case assuming `a ≠ 0` and `a ≠ ∞`, and +`ENNReal.iInf_div` for the special case assuming `Nonempty ι`. -/ +lemma iInf_div' (hinfty : a = 0 → ⨅ i, f i = 0 → ∃ i, f i = 0) (h₀ : a = ∞ → Nonempty ι) : + (⨅ i, f i) / a = ⨅ i, f i / a := iInf_mul' (by simpa) (by simpa) + +/-- If `a ≠ 0` and `a ≠ ∞`, then division by `a` maps infimum to infimum. + +See `ENNReal.iInf_div'` for the general case, and `ENNReal.iInf_div` for another special case that +assumes `Nonempty ι` but does not require `a ≠ ∞`. -/ +lemma iInf_div_of_ne (ha₀ : a ≠ 0) (ha : a ≠ ∞) : (⨅ i, f i) / a = ⨅ i, f i / a := + iInf_div' (by simp [ha₀]) (by simp [ha]) + +/-- See `ENNReal.iInf_div'` for the general case, and `ENNReal.iInf_div_of_ne` for another special +case that assumes `a ≠ ∞` but does not require `Nonempty ι`. -/ +lemma iInf_div [Nonempty ι] (hinfty : a = 0 → ⨅ i, f i = 0 → ∃ i, f i = 0) : + (⨅ i, f i) / a = ⨅ i, f i / a := iInf_div' hinfty fun _ ↦ ‹Nonempty ι› + +lemma inv_iInf (f : ι → ℝ≥0∞) : (⨅ i, f i)⁻¹ = ⨆ i, (f i)⁻¹ := OrderIso.invENNReal.map_iInf _ +lemma inv_iSup (f : ι → ℝ≥0∞) : (⨆ i, f i)⁻¹ = ⨅ i, (f i)⁻¹ := OrderIso.invENNReal.map_iSup _ + +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 add_iSup [Nonempty ι] (f : ι → ℝ≥0∞) : a + ⨆ i, f i = ⨆ i, a + f i := by + obtain rfl | ha := eq_or_ne a ∞ + · simp + refine le_antisymm ?_ <| iSup_le fun i ↦ add_le_add_left (le_iSup ..) _ + refine add_le_of_le_tsub_left_of_le (le_iSup_of_le (Classical.arbitrary _) le_self_add) ?_ + exact iSup_le fun i ↦ ENNReal.le_sub_of_add_le_left ha <| le_iSup (a + f ·) i + +lemma iSup_add [Nonempty ι] (f : ι → ℝ≥0∞) : (⨆ i, f i) + a = ⨆ i, f i + a := by + simp [add_comm, add_iSup] + +lemma add_biSup' {p : ι → Prop} (h : ∃ i, p i) (f : ι → ℝ≥0∞) : + a + ⨆ i, ⨆ _ : p i, f i = ⨆ i, ⨆ _ : p i, a + f i := by + haveI : Nonempty {i // p i} := nonempty_subtype.2 h + simp only [iSup_subtype', add_iSup] + +lemma biSup_add' {p : ι → Prop} (h : ∃ i, p i) (f : ι → ℝ≥0∞) : + (⨆ i, ⨆ _ : p i, f i) + a = ⨆ i, ⨆ _ : p i, f i + a := by simp only [add_comm, add_biSup' h] + +lemma add_biSup {ι : Type*} {s : Set ι} (hs : s.Nonempty) (f : ι → ℝ≥0∞) : + a + ⨆ i ∈ s, f i = ⨆ i ∈ s, a + f i := add_biSup' hs _ + +lemma biSup_add {ι : Type*} {s : Set ι} (hs : s.Nonempty) (f : ι → ℝ≥0∞) : + (⨆ i ∈ s, f i) + a = ⨆ i ∈ s, f i + a := biSup_add' hs _ + +lemma add_sSup (hs : s.Nonempty) : a + sSup s = ⨆ b ∈ s, a + b := by + rw [sSup_eq_iSup, add_biSup hs] + +lemma sSup_add (hs : s.Nonempty) : sSup s + a = ⨆ b ∈ s, b + a := by + rw [sSup_eq_iSup, biSup_add hs] + +lemma iSup_add_iSup_le [Nonempty ι] [Nonempty κ] {g : κ → ℝ≥0∞} (h : ∀ i j, f i + g j ≤ a) : + iSup f + iSup g ≤ a := by simp_rw [iSup_add, add_iSup]; exact iSup₂_le h + +lemma biSup_add_biSup_le' {p : ι → Prop} {q : κ → Prop} (hp : ∃ i, p i) (hq : ∃ j, q j) + {g : κ → ℝ≥0∞} (h : ∀ i, p i → ∀ j, q j → f i + g j ≤ a) : + (⨆ i, ⨆ _ : p i, f i) + ⨆ j, ⨆ _ : q j, g j ≤ a := by + simp_rw [biSup_add' hp, add_biSup' hq] + exact iSup₂_le fun i hi => iSup₂_le (h i hi) + +lemma biSup_add_biSup_le {ι κ : Type*} {s : Set ι} {t : Set κ} (hs : s.Nonempty) (ht : t.Nonempty) + {f : ι → ℝ≥0∞} {g : κ → ℝ≥0∞} {a : ℝ≥0∞} (h : ∀ i ∈ s, ∀ j ∈ t, f i + g j ≤ a) : + (⨆ i ∈ s, f i) + ⨆ j ∈ t, g j ≤ a := biSup_add_biSup_le' hs ht h + +lemma iSup_add_iSup (h : ∀ i j, ∃ k, f i + g j ≤ f k + g k) : iSup f + iSup g = ⨆ i, f i + g i := by + cases isEmpty_or_nonempty ι + · simp only [iSup_of_empty, bot_eq_zero, zero_add] + · refine le_antisymm ?_ (iSup_le fun a => add_le_add (le_iSup _ _) (le_iSup _ _)) + refine iSup_add_iSup_le fun i j => ?_ + rcases h i j with ⟨k, hk⟩ + exact le_iSup_of_le k hk + +lemma iSup_add_iSup_of_monotone {ι : Type*} [Preorder ι] [IsDirected ι (· ≤ ·)] {f g : ι → ℝ≥0∞} + (hf : Monotone f) (hg : Monotone g) : iSup f + iSup g = ⨆ a, f a + g a := + iSup_add_iSup fun i j ↦ (exists_ge_ge i j).imp fun _k ⟨hi, hj⟩ ↦ by gcongr <;> apply_rules + +lemma finsetSum_iSup {α ι : Type*} {s : Finset α} {f : α → ι → ℝ≥0∞} + (hf : ∀ i j, ∃ k, ∀ a, f a i ≤ f a k ∧ f a j ≤ f a k) : + ∑ a ∈ s, ⨆ i, f a i = ⨆ i, ∑ a ∈ s, f a i := by + induction' s using Finset.cons_induction with a s ha ihs + · simp + simp_rw [Finset.sum_cons, ihs] + refine iSup_add_iSup fun i j ↦ (hf i j).imp fun k hk ↦ ?_ + gcongr + exacts [(hk a).1, (hk _).2] + +lemma finsetSum_iSup_of_monotone {α ι : Type*} [Preorder ι] [IsDirected ι (· ≤ ·)] {s : Finset α} + {f : α → ι → ℝ≥0∞} (hf : ∀ a, Monotone (f a)) : (∑ a ∈ s, iSup (f a)) = ⨆ n, ∑ a ∈ s, f a n := + finsetSum_iSup fun i j ↦ (exists_ge_ge i j).imp fun _k ⟨hi, hj⟩ a ↦ ⟨hf a hi, hf a hj⟩ + +@[deprecated (since := "2024-07-14")] +alias finset_sum_iSup_nat := finsetSum_iSup_of_monotone + +lemma smul_iSup {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (f : ι → ℝ≥0∞) (c : R) : + c • ⨆ i, f i = ⨆ i, c • f i := by + simp only [← smul_one_mul c (f _), ← smul_one_mul c (iSup _), ENNReal.mul_iSup] + +lemma smul_sSup {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (s : Set ℝ≥0∞) (c : R) : + c • sSup s = ⨆ a ∈ s, c • a := by + simp_rw [← smul_one_mul c (sSup s), ENNReal.mul_sSup, smul_one_mul] + +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]⟩ + 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) <| + add_le_of_le_tsub_right_of_le (iInf_le_of_le (Classical.arbitrary _) tsub_le_self) <| + iSup_le fun i ↦ ?_ + rw [← sub_sub_cancel ha (h _)] + exact tsub_le_tsub_left (iInf_le (a - f ·) i) _ + +-- TODO: Prove the two one-side versions +lemma exists_lt_add_of_lt_add {x y z : ℝ≥0∞} (h : x < y + z) (hy : y ≠ 0) (hz : z ≠ 0) : + ∃ y' < y, ∃ z' < z, x < y' + z' := by + contrapose! h; + simpa using biSup_add_biSup_le' (by exact ⟨0, hy.bot_lt⟩) (by exact ⟨0, hz.bot_lt⟩) h + end Inv end ENNReal diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 972d93ec15a0c..dbe929eb16e42 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -508,6 +508,19 @@ theorem toReal_sSup (s : Set ℝ≥0∞) (hf : ∀ r ∈ s, r ≠ ∞) : (sSup s).toReal = sSup (ENNReal.toReal '' s) := by simp only [ENNReal.toReal, toNNReal_sSup s hf, NNReal.coe_sSup, Set.image_image] +@[simp] lemma ofReal_iInf [Nonempty ι] (f : ι → ℝ) : + ENNReal.ofReal (⨅ i, f i) = ⨅ i, ENNReal.ofReal (f i) := by + obtain ⟨i, hi⟩ | h := em (∃ i, f i ≤ 0) + · rw [(iInf_eq_bot _).2 fun _ _ ↦ ⟨i, by simpa [ofReal_of_nonpos hi]⟩] + simp [Real.iInf_nonpos' ⟨i, hi⟩] + replace h i : 0 ≤ f i := le_of_not_le fun hi ↦ h ⟨i, hi⟩ + refine eq_of_forall_le_iff fun a ↦ ?_ + obtain rfl | ha := eq_or_ne a ∞ + · simp + rw [le_iInf_iff, le_ofReal_iff_toReal_le ha, le_ciInf_iff ⟨0, by simpa [mem_lowerBounds]⟩] + exact forall_congr' fun i ↦ (le_ofReal_iff_toReal_le ha (h _)).symm + exact Real.iInf_nonneg h + theorem iInf_add : iInf f + a = ⨅ i, f i + a := le_antisymm (le_iInf fun _ => add_le_add (iInf_le _ _) <| le_rfl) (tsub_le_iff_right.1 <| le_iInf fun _ => tsub_le_iff_right.2 <| iInf_le _ _) @@ -545,51 +558,12 @@ theorem iInf_sum {α : Type*} {f : ι → α → ℝ≥0∞} {s : Finset α} [No rw [Finset.forall_mem_cons] at hk exact add_le_add hk.1.1 (Finset.sum_le_sum fun a ha => (hk.2 a ha).2) -/-- If `x ≠ 0` and `x ≠ ∞`, then right multiplication by `x` maps infimum to infimum. -See also `ENNReal.iInf_mul` that assumes `[Nonempty ι]` but does not require `x ≠ 0`. -/ -theorem iInf_mul_of_ne {ι} {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h0 : x ≠ 0) (h : x ≠ ∞) : - iInf f * x = ⨅ i, f i * x := - le_antisymm mul_right_mono.map_iInf_le - ((ENNReal.div_le_iff_le_mul (Or.inl h0) <| Or.inl h).mp <| - le_iInf fun _ => (ENNReal.div_le_iff_le_mul (Or.inl h0) <| Or.inl h).mpr <| iInf_le _ _) - -/-- If `x ≠ ∞`, then right multiplication by `x` maps infimum over a nonempty type to infimum. See -also `ENNReal.iInf_mul_of_ne` that assumes `x ≠ 0` but does not require `[Nonempty ι]`. -/ -theorem iInf_mul {ι} [Nonempty ι] {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h : x ≠ ∞) : - iInf f * x = ⨅ i, f i * x := by - by_cases h0 : x = 0 - · simp only [h0, mul_zero, iInf_const] - · exact iInf_mul_of_ne h0 h - -/-- If `x ≠ ∞`, then left multiplication by `x` maps infimum over a nonempty type to infimum. See -also `ENNReal.mul_iInf_of_ne` that assumes `x ≠ 0` but does not require `[Nonempty ι]`. -/ -theorem mul_iInf {ι} [Nonempty ι] {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h : x ≠ ∞) : - x * iInf f = ⨅ i, x * f i := by simpa only [mul_comm] using iInf_mul h - -/-- If `x ≠ 0` and `x ≠ ∞`, then left multiplication by `x` maps infimum to infimum. -See also `ENNReal.mul_iInf` that assumes `[Nonempty ι]` but does not require `x ≠ 0`. -/ -theorem mul_iInf_of_ne {ι} {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h0 : x ≠ 0) (h : x ≠ ∞) : - x * iInf f = ⨅ i, x * f i := by simpa only [mul_comm] using iInf_mul_of_ne h0 h - -/-! `supr_mul`, `mul_supr` and variants are in `Topology.Instances.ENNReal`. -/ - end iInf section iSup - -@[simp] -theorem iSup_eq_zero {ι : Sort*} {f : ι → ℝ≥0∞} : ⨆ i, f i = 0 ↔ ∀ i, f i = 0 := - iSup_eq_bot - -@[simp] -theorem iSup_zero_eq_zero {ι : Sort*} : ⨆ _ : ι, (0 : ℝ≥0∞) = 0 := by simp - theorem sup_eq_zero {a b : ℝ≥0∞} : a ⊔ b = 0 ↔ a = 0 ∧ b = 0 := sup_eq_bot_iff -theorem iSup_natCast : ⨆ n : ℕ, (n : ℝ≥0∞) = ∞ := - (iSup_eq_top _).2 fun _b hb => ENNReal.exists_nat_gt (lt_top_iff_ne_top.1 hb) - @[deprecated (since := "2024-04-05")] alias iSup_coe_nat := iSup_natCast end iSup diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 4b2e98b074018..9fced5cac44e7 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1995,8 +1995,7 @@ theorem SimpleFunc.exists_lt_lintegral_simpleFunc_of_lt_lintegral {m : Measurabl rcases h₁ hL with ⟨g, g_le, g_top, gL⟩ refine ⟨g, fun x => (g_le x).trans ?_, g_top, gL⟩ simp only [SimpleFunc.coe_add, Pi.add_apply, le_add_iff_nonneg_right, zero_le'] - obtain ⟨L₁, L₂, hL₁, hL₂, hL⟩ : - ∃ L₁ L₂ : ℝ≥0∞, (L₁ < ∫⁻ x, f₁ x ∂μ) ∧ (L₂ < ∫⁻ x, f₂ x ∂μ) ∧ L < L₁ + L₂ := + obtain ⟨L₁, hL₁, L₂, hL₂, hL⟩ : ∃ L₁ < ∫⁻ x, f₁ x ∂μ, ∃ L₂ < ∫⁻ x, f₂ x ∂μ, L < L₁ + L₂ := ENNReal.exists_lt_add_of_lt_add hL hf₁ hf₂ rcases h₁ hL₁ with ⟨g₁, g₁_le, g₁_top, hg₁⟩ rcases h₂ hL₂ with ⟨g₂, g₂_le, g₂_top, hg₂⟩ diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 84bfe07b9e3fb..e03f24092e1a1 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -183,7 +183,7 @@ theorem borel_le_caratheodory (hm : IsMetric μ) : borel X ≤ μ.caratheodory : suffices μ (⋃ n, S n) ≤ ⨆ n, μ (S n) by calc μ (s ∩ t) + μ (s \ t) = μ (s ∩ t) + μ (⋃ n, S n) := by rw [iUnion_S] _ ≤ μ (s ∩ t) + ⨆ n, μ (S n) := by gcongr - _ = ⨆ n, μ (s ∩ t) + μ (S n) := ENNReal.add_iSup + _ = ⨆ n, μ (s ∩ t) + μ (S n) := ENNReal.add_iSup .. _ ≤ μ s := iSup_le hSs /- It suffices to show that `∑' k, μ (S (k + 1) \ S k) ≠ ∞`. Indeed, if we have this, then for all `N` we have `μ (⋃ n, S n) ≤ μ (S N) + ∑' k, m (S (N + k + 1) \ S (N + k))` diff --git a/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean b/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean index a6c90472f6f46..7e859de1147fc 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean @@ -232,7 +232,7 @@ theorem smul_ofFunction {c : ℝ≥0∞} (hc : c ≠ ∞) : c • OuterMeasure.o haveI : Nonempty { t : ℕ → Set α // s ⊆ ⋃ i, t i } := ⟨⟨fun _ => s, subset_iUnion (fun _ => s) 0⟩⟩ simp only [smul_apply, ofFunction_apply, ENNReal.tsum_mul_left, Pi.smul_apply, smul_eq_mul, iInf_subtype'] - rw [ENNReal.iInf_mul_left fun h => (hc h).elim] + rw [ENNReal.mul_iInf fun h => (hc h).elim] end OfFunction diff --git a/Mathlib/Topology/EMetricSpace/Diam.lean b/Mathlib/Topology/EMetricSpace/Diam.lean index 7fb1092cbfdc0..18895f5377d0d 100644 --- a/Mathlib/Topology/EMetricSpace/Diam.lean +++ b/Mathlib/Topology/EMetricSpace/Diam.lean @@ -4,7 +4,6 @@ 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.EMetricSpace.Pi -import Mathlib.Data.ENNReal.Real /-! # Diameters of sets in extended metric spaces diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index fd6a716306a63..609c01f853877 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -435,34 +435,28 @@ theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0∞} (h : ∀ a < 1, a * x ≤ rw [one_mul] at this exact le_of_tendsto this (eventually_nhdsWithin_iff.2 <| Eventually.of_forall h) +@[deprecated mul_iInf' (since := "2024-09-12")] theorem iInf_mul_left' {ι} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) - (h0 : a = 0 → Nonempty ι) : ⨅ i, a * f i = a * ⨅ i, f i := by - by_cases H : a = ∞ ∧ ⨅ i, f i = 0 - · rcases h H.1 H.2 with ⟨i, hi⟩ - rw [H.2, mul_zero, ← bot_eq_zero, iInf_eq_bot] - exact fun b hb => ⟨i, by rwa [hi, mul_zero, ← bot_eq_zero]⟩ - · rw [not_and_or] at H - cases isEmpty_or_nonempty ι - · rw [iInf_of_empty, iInf_of_empty, mul_top] - exact mt h0 (not_nonempty_iff.2 ‹_›) - · exact (ENNReal.mul_left_mono.map_ciInf_of_continuousAt - (ENNReal.continuousAt_const_mul H)).symm + (h0 : a = 0 → Nonempty ι) : ⨅ i, a * f i = a * ⨅ i, f i := .symm <| mul_iInf' h h0 +@[deprecated mul_iInf (since := "2024-09-12")] theorem iInf_mul_left {ι} [Nonempty ι] {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) : ⨅ i, a * f i = a * ⨅ i, f i := - iInf_mul_left' h fun _ => ‹Nonempty ι› + .symm <| mul_iInf h +@[deprecated iInf_mul' (since := "2024-09-12")] theorem iInf_mul_right' {ι} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) - (h0 : a = 0 → Nonempty ι) : ⨅ i, f i * a = (⨅ i, f i) * a := by - simpa only [mul_comm a] using iInf_mul_left' h h0 + (h0 : a = 0 → Nonempty ι) : ⨅ i, f i * a = (⨅ i, f i) * a := .symm <| iInf_mul' h h0 +@[deprecated iInf_mul (since := "2024-09-12")] theorem iInf_mul_right {ι} [Nonempty ι] {f : ι → ℝ≥0∞} {a : ℝ≥0∞} - (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) : ⨅ i, f i * a = (⨅ i, f i) * a := - iInf_mul_right' h fun _ => ‹Nonempty ι› + (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) : ⨅ i, f i * a = (⨅ i, f i) * a := .symm <| iInf_mul h +@[deprecated inv_iInf (since := "2024-09-12")] theorem inv_map_iInf {ι : Sort*} {x : ι → ℝ≥0∞} : (iInf x)⁻¹ = ⨆ i, (x i)⁻¹ := OrderIso.invENNReal.map_iInf x +@[deprecated inv_iSup (since := "2024-09-12")] theorem inv_map_iSup {ι : Sort*} {x : ι → ℝ≥0∞} : (iSup x)⁻¹ = ⨅ i, (x i)⁻¹ := OrderIso.invENNReal.map_iSup x @@ -504,115 +498,10 @@ protected theorem Tendsto.div_const {f : Filter α} {m : α → ℝ≥0∞} {a b protected theorem tendsto_inv_nat_nhds_zero : Tendsto (fun n : ℕ => (n : ℝ≥0∞)⁻¹) atTop (𝓝 0) := ENNReal.inv_top ▸ ENNReal.tendsto_inv_iff.2 tendsto_nat_nhds_top -theorem iSup_add {ι : Sort*} {s : ι → ℝ≥0∞} [Nonempty ι] : iSup s + a = ⨆ b, s b + a := - Monotone.map_ciSup_of_continuousAt (continuousAt_id.add continuousAt_const) <| - monotone_id.add monotone_const - -theorem biSup_add' {ι : Sort*} {p : ι → Prop} (h : ∃ i, p i) {f : ι → ℝ≥0∞} : - (⨆ (i) (_ : p i), f i) + a = ⨆ (i) (_ : p i), f i + a := by - haveI : Nonempty { i // p i } := nonempty_subtype.2 h - simp only [iSup_subtype', iSup_add] - -theorem add_biSup' {ι : Sort*} {p : ι → Prop} (h : ∃ i, p i) {f : ι → ℝ≥0∞} : - (a + ⨆ (i) (_ : p i), f i) = ⨆ (i) (_ : p i), a + f i := by - simp only [add_comm a, biSup_add' h] - -theorem biSup_add {ι} {s : Set ι} (hs : s.Nonempty) {f : ι → ℝ≥0∞} : - (⨆ i ∈ s, f i) + a = ⨆ i ∈ s, f i + a := - biSup_add' hs - -theorem add_biSup {ι} {s : Set ι} (hs : s.Nonempty) {f : ι → ℝ≥0∞} : - (a + ⨆ i ∈ s, f i) = ⨆ i ∈ s, a + f i := - add_biSup' hs - -theorem sSup_add {s : Set ℝ≥0∞} (hs : s.Nonempty) : sSup s + a = ⨆ b ∈ s, b + a := by - rw [sSup_eq_iSup, biSup_add hs] - -theorem add_iSup {ι : Sort*} {s : ι → ℝ≥0∞} [Nonempty ι] : a + iSup s = ⨆ b, a + s b := by - rw [add_comm, iSup_add]; simp [add_comm] - -theorem iSup_add_iSup_le {ι ι' : Sort*} [Nonempty ι] [Nonempty ι'] {f : ι → ℝ≥0∞} {g : ι' → ℝ≥0∞} - {a : ℝ≥0∞} (h : ∀ i j, f i + g j ≤ a) : iSup f + iSup g ≤ a := by - simp_rw [iSup_add, add_iSup]; exact iSup₂_le h - -theorem biSup_add_biSup_le' {ι ι'} {p : ι → Prop} {q : ι' → Prop} (hp : ∃ i, p i) (hq : ∃ j, q j) - {f : ι → ℝ≥0∞} {g : ι' → ℝ≥0∞} {a : ℝ≥0∞} (h : ∀ i, p i → ∀ j, q j → f i + g j ≤ a) : - ((⨆ (i) (_ : p i), f i) + ⨆ (j) (_ : q j), g j) ≤ a := by - simp_rw [biSup_add' hp, add_biSup' hq] - exact iSup₂_le fun i hi => iSup₂_le (h i hi) - -theorem biSup_add_biSup_le {ι ι'} {s : Set ι} {t : Set ι'} (hs : s.Nonempty) (ht : t.Nonempty) - {f : ι → ℝ≥0∞} {g : ι' → ℝ≥0∞} {a : ℝ≥0∞} (h : ∀ i ∈ s, ∀ j ∈ t, f i + g j ≤ a) : - ((⨆ i ∈ s, f i) + ⨆ j ∈ t, g j) ≤ a := - biSup_add_biSup_le' hs ht h - -theorem iSup_add_iSup {ι : Sort*} {f g : ι → ℝ≥0∞} (h : ∀ i j, ∃ k, f i + g j ≤ f k + g k) : - iSup f + iSup g = ⨆ a, f a + g a := by - cases isEmpty_or_nonempty ι - · simp only [iSup_of_empty, bot_eq_zero, zero_add] - · refine le_antisymm ?_ (iSup_le fun a => add_le_add (le_iSup _ _) (le_iSup _ _)) - refine iSup_add_iSup_le fun i j => ?_ - rcases h i j with ⟨k, hk⟩ - exact le_iSup_of_le k hk - -theorem iSup_add_iSup_of_monotone {ι : Type*} [Preorder ι] [IsDirected ι (· ≤ ·)] - {f g : ι → ℝ≥0∞} (hf : Monotone f) (hg : Monotone g) : iSup f + iSup g = ⨆ a, f a + g a := - iSup_add_iSup fun i j ↦ (exists_ge_ge i j).imp fun _k ⟨hi, hj⟩ ↦ by gcongr <;> apply_rules - -theorem finsetSum_iSup {α ι : Type*} {s : Finset α} {f : α → ι → ℝ≥0∞} - (hf : ∀ i j, ∃ k, ∀ a, f a i ≤ f a k ∧ f a j ≤ f a k) : - ∑ a ∈ s, ⨆ i, f a i = ⨆ i, ∑ a ∈ s, f a i := by - induction s using Finset.cons_induction with - | empty => simp - | cons a s ha ihs => - simp_rw [Finset.sum_cons, ihs] - refine iSup_add_iSup fun i j ↦ (hf i j).imp fun k hk ↦ ?_ - gcongr - exacts [(hk a).1, (hk _).2] - -theorem finsetSum_iSup_of_monotone {α} {ι} [Preorder ι] [IsDirected ι (· ≤ ·)] - {s : Finset α} {f : α → ι → ℝ≥0∞} (hf : ∀ a, Monotone (f a)) : - (∑ a ∈ s, iSup (f a)) = ⨆ n, ∑ a ∈ s, f a n := - finsetSum_iSup fun i j ↦ (exists_ge_ge i j).imp fun _k ⟨hi, hj⟩ a ↦ ⟨hf a hi, hf a hj⟩ - -@[deprecated (since := "2024-07-14")] -alias finset_sum_iSup_nat := finsetSum_iSup_of_monotone - -theorem mul_iSup {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : a * iSup f = ⨆ i, a * f i := by - by_cases hf : ∀ i, f i = 0 - · obtain rfl : f = fun _ => 0 := funext hf - simp only [iSup_zero_eq_zero, mul_zero] - · refine (monotone_id.const_mul' _).map_iSup_of_continuousAt ?_ (mul_zero a) - refine ENNReal.Tendsto.const_mul tendsto_id (Or.inl ?_) - exact mt iSup_eq_zero.1 hf - -theorem mul_sSup {s : Set ℝ≥0∞} {a : ℝ≥0∞} : a * sSup s = ⨆ i ∈ s, a * i := by - simp only [sSup_eq_iSup, mul_iSup] - -theorem iSup_mul {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : iSup f * a = ⨆ i, f i * a := by - rw [mul_comm, mul_iSup]; congr; funext; rw [mul_comm] - -theorem smul_iSup {ι : Sort*} {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (f : ι → ℝ≥0∞) - (c : R) : (c • ⨆ i, f i) = ⨆ i, c • f i := by - -- Porting note: replaced `iSup _` with `iSup f` - simp only [← smul_one_mul c (f _), ← smul_one_mul c (iSup f), ENNReal.mul_iSup] - -theorem smul_sSup {R} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (s : Set ℝ≥0∞) (c : R) : - c • sSup s = ⨆ i ∈ s, c • i := by - -- Porting note: replaced `_` with `s` - simp_rw [← smul_one_mul c (sSup s), ENNReal.mul_sSup, smul_one_mul] - -theorem iSup_div {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : iSup f / a = ⨆ i, f i / a := - iSup_mul - protected theorem tendsto_coe_sub {b : ℝ≥0∞} : Tendsto (fun b : ℝ≥0∞ => ↑r - b) (𝓝 b) (𝓝 (↑r - b)) := continuous_nnreal_sub.tendsto _ -theorem sub_iSup {ι : Sort*} [Nonempty ι] {b : ι → ℝ≥0∞} (hr : a < ∞) : - (a - ⨆ i, b i) = ⨅ i, a - b i := - antitone_const_tsub.map_ciSup_of_continuousAt (continuous_sub_left hr.ne).continuousAt - theorem exists_countable_dense_no_zero_top : ∃ s : Set ℝ≥0∞, s.Countable ∧ Dense s ∧ 0 ∉ s ∧ ∞ ∉ s := by obtain ⟨s, s_count, s_dense, hs⟩ : @@ -620,19 +509,7 @@ theorem exists_countable_dense_no_zero_top : exists_countable_dense_no_bot_top ℝ≥0∞ exact ⟨s, s_count, s_dense, fun h => hs.1 0 (by simp) h, fun h => hs.2 ∞ (by simp) h⟩ -theorem exists_lt_add_of_lt_add {x y z : ℝ≥0∞} (h : x < y + z) (hy : y ≠ 0) (hz : z ≠ 0) : - ∃ y' z', y' < y ∧ z' < z ∧ x < y' + z' := by - have : NeZero y := ⟨hy⟩ - have : NeZero z := ⟨hz⟩ - have A : Tendsto (fun p : ℝ≥0∞ × ℝ≥0∞ => p.1 + p.2) (𝓝[<] y ×ˢ 𝓝[<] z) (𝓝 (y + z)) := by - apply Tendsto.mono_left _ (Filter.prod_mono nhdsWithin_le_nhds nhdsWithin_le_nhds) - rw [← nhds_prod_eq] - exact tendsto_add - rcases ((A.eventually (lt_mem_nhds h)).and - (Filter.prod_mem_prod self_mem_nhdsWithin self_mem_nhdsWithin)).exists with - ⟨⟨y', z'⟩, hx, hy', hz'⟩ - exact ⟨y', z', hy', hz', hx⟩ - +@[deprecated ofReal_iInf (since := "2024-09-12")] theorem ofReal_cinfi (f : α → ℝ) [Nonempty α] : ENNReal.ofReal (⨅ i, f i) = ⨅ i, ENNReal.ofReal (f i) := by by_cases hf : BddBelow (range f) @@ -1526,5 +1403,3 @@ lemma limsup_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} ( end LimsupLiminf end ENNReal -- namespace - -set_option linter.style.longFile 1700 From 20a854cda8bedffea00955e1b174ae4c7f72b269 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Oct 2024 21:25:32 +0000 Subject: [PATCH 139/174] feat: `star` commutes with `nnqsmul` (#17351) From LeanAPAP --- Mathlib/Algebra/Star/Module.lean | 38 +++++++++++++++++++++++++++++--- Mathlib/Data/Matrix/Basic.lean | 2 -- 2 files changed, 35 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Star/Module.lean b/Mathlib/Algebra/Star/Module.lean index 1640c1f29979a..214d513fed01a 100644 --- a/Mathlib/Algebra/Star/Module.lean +++ b/Mathlib/Algebra/Star/Module.lean @@ -75,11 +75,43 @@ theorem star_ratCast_smul [DivisionRing R] [AddCommGroup M] [Module R M] [StarAd @[deprecated (since := "2024-04-17")] alias star_rat_cast_smul := star_ratCast_smul -@[simp] -theorem star_rat_smul {R : Type*} [AddCommGroup R] [StarAddMonoid R] [Module ℚ R] (x : R) (n : ℚ) : - star (n • x) = n • star x := +/-! +Per the naming convention, these two lemmas call `(q • ·)` `nnrat_smul` and `rat_smul` respectively, +rather than `nnqsmul` and `qsmul` because the latter are reserved to the actions coming from +`DivisionSemiring` and `DivisionRing`. We provide aliases with `nnqsmul` and `qsmul` for +discoverability. +-/ + +/-- Note that this lemma holds for an arbitrary `ℚ≥0`-action, rather than merely one coming from a +`DivisionSemiring`. We keep both the `nnqsmul` and `nnrat_smul` naming conventions for +discoverability. See `star_nnqsmul`. -/ +@[simp high] +lemma star_nnrat_smul [AddCommMonoid R] [StarAddMonoid R] [Module ℚ≥0 R] (q : ℚ≥0) (x : R) : + star (q • x) = q • star x := map_nnrat_smul (starAddEquiv : R ≃+ R) _ _ + +/-- Note that this lemma holds for an arbitrary `ℚ`-action, rather than merely one coming from a +`DivisionRing`. We keep both the `qsmul` and `rat_smul` naming conventions for discoverability. +See `star_qsmul`. -/ +@[simp high] lemma star_rat_smul [AddCommGroup R] [StarAddMonoid R] [Module ℚ R] (q : ℚ) (x : R) : + star (q • x) = q • star x := map_rat_smul (starAddEquiv : R ≃+ R) _ _ +/-- Note that this lemma holds for an arbitrary `ℚ≥0`-action, rather than merely one coming from a +`DivisionSemiring`. We keep both the `nnqsmul` and `nnrat_smul` naming conventions for +discoverability. See `star_nnrat_smul`. -/ +alias star_nnqsmul := star_nnrat_smul + +/-- Note that this lemma holds for an arbitrary `ℚ`-action, rather than merely one coming from a +`DivisionRing`. We keep both the `qsmul` and `rat_smul` naming conventions for +discoverability. See `star_rat_smul`. -/ +alias star_qsmul := star_rat_smul + +instance StarAddMonoid.toStarModuleNNRat [AddCommMonoid R] [Module ℚ≥0 R] [StarAddMonoid R] : + StarModule ℚ≥0 R where star_smul := star_nnrat_smul + +instance StarAddMonoid.toStarModuleRat [AddCommGroup R] [Module ℚ R] [StarAddMonoid R] : + StarModule ℚ R where star_smul := star_rat_smul + end SMulLemmas /-- If `A` is a module over a commutative `R` with compatible actions, diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 6bcfb3a9df82c..63bc543ba910b 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -2098,7 +2098,6 @@ variants which this lemma would not apply to: * `Matrix.conjTranspose_intCast_smul` * `Matrix.conjTranspose_inv_natCast_smul` * `Matrix.conjTranspose_inv_intCast_smul` -* `Matrix.conjTranspose_rat_smul` * `Matrix.conjTranspose_ratCast_smul` -/ @[simp] @@ -2166,7 +2165,6 @@ theorem conjTranspose_ratCast_smul [DivisionRing R] [AddCommGroup α] [StarAddMo (c : ℚ) (M : Matrix m n α) : ((c : R) • M)ᴴ = (c : R) • Mᴴ := Matrix.ext <| by simp -@[simp] theorem conjTranspose_rat_smul [AddCommGroup α] [StarAddMonoid α] [Module ℚ α] (c : ℚ) (M : Matrix m n α) : (c • M)ᴴ = c • Mᴴ := Matrix.ext <| by simp From 655d653bfd6eb4968829ce8e2b5e4ed0419f6f05 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Wed, 2 Oct 2024 22:35:15 +0000 Subject: [PATCH 140/174] feat (LinearAlgebra/RootPairing/Finite) : The canonical bilinear form on a finite root pairing (#17070) This PR introduces the canonical bilinear form for a finite root pairing, and proves some basic properties. --- Mathlib.lean | 1 + Mathlib/Algebra/Ring/SumsOfSquares.lean | 12 +- Mathlib/LinearAlgebra/RootSystem/Basic.lean | 4 +- Mathlib/LinearAlgebra/RootSystem/Defs.lean | 62 ++++++-- .../RootSystem/Finite/CanonicalBilinear.lean | 137 ++++++++++++++++++ Mathlib/LinearAlgebra/RootSystem/Hom.lean | 4 +- .../RootSystem/RootPositive.lean | 2 +- 7 files changed, 207 insertions(+), 15 deletions(-) create mode 100644 Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0d6d6971bf5da..128adb94db271 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3100,6 +3100,7 @@ import Mathlib.LinearAlgebra.Ray import Mathlib.LinearAlgebra.Reflection import Mathlib.LinearAlgebra.RootSystem.Basic import Mathlib.LinearAlgebra.RootSystem.Defs +import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear import Mathlib.LinearAlgebra.RootSystem.Hom import Mathlib.LinearAlgebra.RootSystem.OfBilinear import Mathlib.LinearAlgebra.RootSystem.RootPairingCat diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean index 1eb9bb77744cf..6abec158dd334 100644 --- a/Mathlib/Algebra/Ring/SumsOfSquares.lean +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -3,9 +3,8 @@ Copyright (c) 2024 Florent Schaffhauser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Florent Schaffhauser -/ -import Mathlib.Algebra.Ring.Defs +import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Submonoid.Basic -import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Order.Ring.Defs /-! @@ -61,6 +60,15 @@ theorem IsSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : IsSumSq S1) @[deprecated (since := "2024-08-09")] alias isSumSq.add := IsSumSq.add +/-- A finite sum of squares is a sum of squares. -/ +theorem isSumSq_sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : ι → R) : + IsSumSq (∑ i ∈ s, f i * f i) := by + induction s using Finset.cons_induction with + | empty => + simpa only [Finset.sum_empty] using IsSumSq.zero + | cons i s his h => + exact (Finset.sum_cons (β := R) his) ▸ IsSumSq.sq_add (f i) (∑ i ∈ s, f i * f i) h + variable (R) in /-- In an additive monoid with multiplication `R`, the type `sumSqIn R` is the submonoid of sums of diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index fd07952c7174a..4aba11c944d5e 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -120,8 +120,8 @@ protected lemma ext [CharZero R] [NoZeroSMulDivisors R M] ext i j refine P₁.root.injective ?_ conv_rhs => rw [hr] - rw [root_reflection_perm, root_reflection_perm] - simp only [hr, he, hc', reflection_apply] + simp only [root_reflection_perm, reflection_apply, coroot'] + simp only [hr, he, hc'] suffices P₁.coroot = P₂.coroot by cases' P₁ with p₁; cases' P₂ with p₂; cases p₁; cases p₂; congr; exact hp this have := NoZeroSMulDivisors.int_of_charZero R M diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 55112f1e275ed..607057bd4e391 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -32,7 +32,6 @@ This file contains basic definitions for root systems and root data. ## TODO * Base change of root pairings (may need flatness; perhaps should go in a different file). - * Isomorphism of root pairings. * Crystallographic root systems are isomorphic to base changes of root systems over `ℤ`: Take `M₀` and `N₀` to be the `ℤ`-span of roots and coroots. @@ -51,7 +50,6 @@ between roots and coroots is (implicitly) included and the coroots are included Empirically this seems to be by far the most convenient design and by providing extensionality lemmas expressing the uniqueness we expect to get the best of both worlds. - Furthermore, we require roots and coroots to be injections from a base indexing type `ι` rather than subsets of their codomains. This design was chosen to avoid the bijection between roots and coroots being a dependently-typed object. A third option would be to have the roots and coroots be subsets @@ -146,13 +144,27 @@ protected def flip : RootPairing ι R N M := lemma flip_flip : P.flip.flip = P := rfl +/-- Roots written as functionals on the coweight space. -/ +abbrev root' (i : ι) : Dual R N := P.toPerfectPairing (P.root i) + +/-- Coroots written as functionals on the weight space. -/ +abbrev coroot' (i : ι) : Dual R M := P.toPerfectPairing.flip (P.coroot i) + /-- This is the pairing between roots and coroots. -/ -def pairing : R := P.toPerfectPairing (P.root i) (P.coroot j) +def pairing : R := P.root' i (P.coroot j) @[simp] lemma root_coroot_eq_pairing : P.toPerfectPairing (P.root i) (P.coroot j) = P.pairing i j := rfl +@[simp] +lemma root'_coroot_eq_pairing : P.root' i (P.coroot j) = P.pairing i j := + rfl + +@[simp] +lemma root_coroot'_eq_pairing : P.coroot' i (P.root j) = P.pairing j i := + rfl + lemma coroot_root_eq_pairing : P.toLin.flip (P.coroot i) (P.root j) = P.pairing j i := by simp @@ -178,7 +190,7 @@ theorem mapsTo_reflection_root : exact P.root_reflection_perm i j ▸ mem_range_self (P.reflection_perm i j) lemma reflection_apply (x : M) : - P.reflection i x = x - (P.toPerfectPairing x (P.coroot i)) • P.root i := + P.reflection i x = x - (P.coroot' i x) • P.root i := rfl lemma reflection_apply_root : @@ -253,7 +265,7 @@ theorem mapsTo_coreflection_coroot : exact P.coroot_reflection_perm i j ▸ mem_range_self (P.reflection_perm i j) lemma coreflection_apply (f : N) : - P.coreflection i f = f - (P.toPerfectPairing (P.root i) f) • P.coroot i := + P.coreflection i f = f - (P.root' i) f • P.coroot i := rfl lemma coreflection_apply_coroot : @@ -303,10 +315,42 @@ lemma coroot_eq_coreflection_of_root_eq rw [← P.root_reflection_perm, EmbeddingLike.apply_eq_iff_eq] at hk rw [← P.coroot_reflection_perm, hk] +lemma coroot'_reflection_perm {i j : ι} : + P.coroot' (P.reflection_perm i j) = P.coroot' j ∘ₗ P.reflection i := by + ext y + simp [coreflection_apply_coroot, reflection_apply, map_sub, mul_comm] + +lemma coroot'_reflection {i j : ι} (y : M) : + P.coroot' j (P.reflection i y) = P.coroot' (P.reflection_perm i j) y := + (LinearMap.congr_fun P.coroot'_reflection_perm y).symm + +lemma pairing_reflection_perm (i j k : ι) : + P.pairing j (P.reflection_perm i k) = P.pairing (P.reflection_perm i j) k := by + simp only [pairing, root', coroot_reflection_perm, root_reflection_perm] + simp only [coreflection_apply_coroot, map_sub, map_smul, smul_eq_mul, + reflection_apply_root] + simp only [← toLin_toPerfectPairing, map_smul, LinearMap.smul_apply, map_sub, map_smul, + LinearMap.sub_apply, smul_eq_mul] + simp only [PerfectPairing.toLin_apply, root'_coroot_eq_pairing, sub_right_inj, mul_comm] + +@[simp] +lemma pairing_reflection_perm_self_left (P : RootPairing ι R M N) (i j : ι) : + P.pairing (P.reflection_perm i i) j = - P.pairing i j := by + rw [pairing, root', ← reflection_perm_root, root'_coroot_eq_pairing, pairing_same, two_smul, + sub_add_cancel_left, ← toLin_toPerfectPairing, LinearMap.map_neg₂, toLin_toPerfectPairing, + root'_coroot_eq_pairing] + +@[simp] +lemma pairing_reflection_perm_self_right (i j : ι) : + P.pairing i (P.reflection_perm j j) = - P.pairing i j := by + rw [pairing, ← reflection_perm_coroot, root_coroot_eq_pairing, pairing_same, two_smul, + sub_add_cancel_left, ← toLin_toPerfectPairing, map_neg, toLin_toPerfectPairing, + root_coroot_eq_pairing] + /-- A root pairing is said to be crystallographic if the pairing between a root and coroot is always an integer. -/ def IsCrystallographic : Prop := - ∀ i, MapsTo (P.toPerfectPairing (P.root i)) (range P.coroot) (zmultiples (1 : R)) + ∀ i, MapsTo (P.root' i) (range P.coroot) (zmultiples (1 : R)) lemma isCrystallographic_iff : P.IsCrystallographic ↔ ∀ i j, ∃ z : ℤ, z = P.pairing i j := by @@ -488,12 +532,12 @@ lemma IsOrthogonal.symm : IsOrthogonal P i j ↔ IsOrthogonal P j i := by simp only [IsOrthogonal, and_comm] lemma isOrthogonal_comm (h : IsOrthogonal P i j) : Commute (P.reflection i) (P.reflection j) := by - rw [Commute, SemiconjBy] + rw [commute_iff_eq] ext v replace h : P.pairing i j = 0 ∧ P.pairing j i = 0 := by simpa [IsOrthogonal] using h erw [LinearMap.mul_apply, LinearMap.mul_apply] - simp only [LinearEquiv.coe_coe, reflection_apply, map_sub, map_smul, root_coroot_eq_pairing, - zero_smul, sub_zero, toLin_toPerfectPairing, h] + simp only [LinearEquiv.coe_coe, reflection_apply, PerfectPairing.flip_apply_apply, map_sub, + map_smul, root_coroot_eq_pairing, h, zero_smul, sub_zero] abel end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean new file mode 100644 index 0000000000000..e7a7f1744d358 --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -0,0 +1,137 @@ +/- +Copyright (c) 2024 Scott Carnahan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Carnahan +-/ +import Mathlib.LinearAlgebra.RootSystem.Defs +import Mathlib.Algebra.Ring.SumsOfSquares + +/-! +# The canonical bilinear form on a finite root pairing +Given a finite root pairing, we define a canonical map from weight space to coweight space, and the +corresponding bilinear form. This form is symmetric and Weyl-invariant, and if the base ring is +linearly ordered, then the form is root-positive, positive-semidefinite on the weight space, and +positive-definite on the span of roots. +From these facts, it is easy to show that Coxeter weights in a finite root pairing are bounded +above by 4. Thus, the pairings of roots and coroots in a crystallographic root pairing are +restricted to a small finite set of possibilities. +Another application is to the faithfulness of the Weyl group action on roots, and finiteness of the +Weyl group. + +## Main definitions: + * `Polarization`: A distinguished linear map from the weight space to the coweight space. + * `RootForm` : The bilinear form on weight space corresponding to `Polarization`. + +## References: + * SGAIII Exp. XXI + * Bourbaki, Lie groups and Lie algebras + +## Main results: + * `polarization_self_sum_of_squares` : The inner product of any weight vector is a sum of squares. + * `rootForm_reflection_reflection_apply` : `RootForm` is invariant with respect + to reflections. + * `rootForm_self_smul_coroot`: The inner product of a root with itself times the + corresponding coroot is equal to two times Polarization applied to the root. + +## TODO (possibly in other files) + * Positivity and nondegeneracy + * Weyl-invariance + * Faithfulness of Weyl group action, and finiteness of Weyl group, for finite root systems. + * Relation to Coxeter weight. In particular, positivity constraints for finite root pairings mean + we restrict to weights between 0 and 4. +-/ + +open Function +open Module hiding reflection + +noncomputable section + +variable {ι R M N : Type*} + +namespace RootPairing + +section CommRing + +variable [Fintype ι] [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] +(P : RootPairing ι R M N) + +/-- An invariant linear map from weight space to coweight space. -/ +def Polarization : M →ₗ[R] N := + ∑ i, LinearMap.toSpanSingleton R N (P.coroot i) ∘ₗ P.coroot' i + +@[simp] +lemma Polarization_apply (x : M) : + P.Polarization x = ∑ i, P.coroot' i x • P.coroot i := by + simp [Polarization] + +/-- An invariant linear map from coweight space to weight space. -/ +def CoPolarization : N →ₗ[R] M := + ∑ i, LinearMap.toSpanSingleton R M (P.root i) ∘ₗ P.root' i + +@[simp] +lemma CoPolarization_apply (x : N) : + P.CoPolarization x = ∑ i, P.root' i x • P.root i := by + simp [CoPolarization] + +lemma CoPolarization_eq : P.CoPolarization = P.flip.Polarization := + rfl + +/-- An invariant inner product on the weight space. -/ +def RootForm : LinearMap.BilinForm R M := + ∑ i, (LinearMap.lsmul R R).compl₁₂ (P.coroot' i) (P.coroot' i) + +/-- An invariant inner product on the coweight space. -/ +def CorootForm : LinearMap.BilinForm R N := + ∑ i, (LinearMap.lsmul R R).compl₁₂ (P.root' i) (P.root' i) + +lemma rootForm_apply_apply (x y : M) : P.RootForm x y = + ∑ (i : ι), P.coroot' i x * P.coroot' i y := by + simp [RootForm] + +lemma rootForm_symmetric : + LinearMap.IsSymm P.RootForm := by + simp [LinearMap.IsSymm, mul_comm, rootForm_apply_apply] + +@[simp] +lemma rootForm_reflection_reflection_apply (i : ι) (x y : M) : + P.RootForm (P.reflection i x) (P.reflection i y) = P.RootForm x y := by + simp only [rootForm_apply_apply, coroot'_reflection] + exact Fintype.sum_equiv (P.reflection_perm i) + (fun j ↦ (P.coroot' (P.reflection_perm i j) x) * (P.coroot' (P.reflection_perm i j) y)) + (fun j ↦ P.coroot' j x * P.coroot' j y) (congrFun rfl) + +/-- This is SGA3 XXI Lemma 1.2.1 (10), key for proving nondegeneracy and positivity. -/ +lemma rootForm_self_smul_coroot (i : ι) : + (P.RootForm (P.root i) (P.root i)) • P.coroot i = 2 • P.Polarization (P.root i) := by + have hP : P.Polarization (P.root i) = + ∑ j : ι, P.pairing i (P.reflection_perm i j) • P.coroot (P.reflection_perm i j) := by + simp_rw [Polarization_apply, root_coroot'_eq_pairing] + exact (Fintype.sum_equiv (P.reflection_perm i) + (fun j ↦ P.pairing i (P.reflection_perm i j) • P.coroot (P.reflection_perm i j)) + (fun j ↦ P.pairing i j • P.coroot j) (congrFun rfl)).symm + rw [two_nsmul] + nth_rw 2 [hP] + rw [Polarization_apply] + simp only [root_coroot'_eq_pairing, pairing_reflection_perm, pairing_reflection_perm_self_left, + ← reflection_perm_coroot, smul_sub, neg_smul, sub_neg_eq_add] + rw [Finset.sum_add_distrib, ← add_assoc, ← sub_eq_iff_eq_add] + simp only [rootForm_apply_apply, LinearMap.coe_comp, comp_apply, Polarization_apply, + root_coroot_eq_pairing, map_sum, LinearMapClass.map_smul, Finset.sum_neg_distrib, ← smul_assoc] + rw [Finset.sum_smul, add_neg_eq_zero.mpr rfl] + exact sub_eq_zero_of_eq rfl + +lemma corootForm_self_smul_root (i : ι) : + (P.CorootForm (P.coroot i) (P.coroot i)) • P.root i = 2 • P.CoPolarization (P.coroot i) := + rootForm_self_smul_coroot (P.flip) i + +lemma rootForm_self_sum_of_squares (x : M) : + IsSumSq (P.RootForm x x) := + P.rootForm_apply_apply x x ▸ isSumSq_sum_mul_self Finset.univ _ + +lemma rootForm_root_self (j : ι) : + P.RootForm (P.root j) (P.root j) = ∑ (i : ι), (P.pairing j i) * (P.pairing j i) := by + simp [rootForm_apply_apply] + +end CommRing + +end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Hom.lean b/Mathlib/LinearAlgebra/RootSystem/Hom.lean index f70d0830fdaee..5d58b28dabbf4 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Hom.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Hom.lean @@ -116,7 +116,9 @@ def Hom.reflection (P : RootPairing ι R M N) (i : ι) : Hom P P where PerfectPairing.toDualRight_apply, LinearMap.dualMap_apply, PerfectPairing.flip_apply_apply, LinearEquiv.comp_coe, LinearEquiv.trans_apply] rw [RootPairing.reflection_apply, RootPairing.coreflection_apply] - simp [map_sub, ← PerfectPairing.toLin_apply, mul_comm] + simp only [← PerfectPairing.toLin_apply, map_sub, map_smul, LinearMap.sub_apply, + toLin_toPerfectPairing, LinearMap.smul_apply, smul_eq_mul, sub_right_inj] + simp only [PerfectPairing.toLin_apply, PerfectPairing.flip_apply_apply, mul_comm] root_weightMap := by ext; simp coroot_coweightMap := by ext; simp diff --git a/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean b/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean index ccf50330e9a81..e9d55f92f8209 100644 --- a/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean +++ b/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean @@ -55,7 +55,7 @@ lemma two_mul_apply_root_root : 2 * B (P.root i) (P.root j) = P.pairing i j * B (P.root j) (P.root j) := by rw [two_mul, ← eq_sub_iff_add_eq] nth_rw 1 [← IsRootPositive.apply_reflection_eq (P := P) (B := B) j (P.root i) (P.root j)] - rw [reflection_apply, reflection_apply_self, root_coroot_eq_pairing, LinearMap.map_sub₂, + rw [reflection_apply, reflection_apply_self, root_coroot'_eq_pairing, LinearMap.map_sub₂, LinearMap.map_smul₂, smul_eq_mul, LinearMap.map_neg, LinearMap.map_neg, mul_neg, neg_sub_neg] @[simp] From f7338d0d29be1b1e472b816d34454d377c8549bd Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 3 Oct 2024 04:46:10 +0000 Subject: [PATCH 141/174] fix: docPrime unlint private declarations (#17361) Mentioned in [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/theorem.20names.20with.20primes.3F/near/474282217). --- Mathlib/Tactic/Linter/DocPrime.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Tactic/Linter/DocPrime.lean b/Mathlib/Tactic/Linter/DocPrime.lean index 86b12eaad3719..79a08666bae3f 100644 --- a/Mathlib/Tactic/Linter/DocPrime.lean +++ b/Mathlib/Tactic/Linter/DocPrime.lean @@ -40,6 +40,8 @@ def docPrimeLinter : Linter where run := withSetOptionIn fun stx ↦ do if (← get).messages.hasErrors then return unless [``Lean.Parser.Command.declaration, `lemma].contains stx.getKind do return + -- ignore private declarations + if (stx.find? (·.isOfKind ``Lean.Parser.Command.private)).isSome then return let docstring := stx[0][0] -- The current declaration's id, possibly followed by a list of universe names. let declId := From 491935d5063cca45cb10602fbecc29a9aaf41700 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 05:41:52 +0000 Subject: [PATCH 142/174] chore: unify monotonicity of rank lemmas (#17292) The names were all over the place. Rename: * `rank_le_of_submodule` -> `Submodule.rank_mono` * `Submodule.finrank_le_finrank_of_le` -> `Submodule.finrank_mono` (existing, up to unfolding `Monotone`) * `rank_submodule_le` -> `Submodule.rank_le` From LeanCamCombi --- Archive/Sensitivity.lean | 2 +- .../AffineSpace/FiniteDimensional.lean | 6 +++--- Mathlib/LinearAlgebra/Dimension/Basic.lean | 11 ++++++---- .../Dimension/Constructions.lean | 12 +++++------ .../LinearAlgebra/Dimension/LinearMap.lean | 6 +++--- .../LinearAlgebra/FiniteDimensional/Defs.lean | 20 ++----------------- Mathlib/LinearAlgebra/LinearDisjoint.lean | 2 +- 7 files changed, 23 insertions(+), 36 deletions(-) diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index 3ebec47f3b7d6..d28ffea139084 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -374,7 +374,7 @@ theorem exists_eigenvalue (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) : suffices 0 < dim (W ⊓ img) by exact mod_cast exists_mem_ne_zero_of_rank_pos this have dim_le : dim (W ⊔ img) ≤ 2 ^ (m + 1 : Cardinal) := by - convert ← rank_submodule_le (W ⊔ img) + convert ← Submodule.rank_le (W ⊔ img) rw [← Nat.cast_succ] apply dim_V have dim_add : dim (W ⊔ img) + dim (W ⊓ img) = dim W + 2 ^ m := by diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index e59a26e0c1ce5..e0cc46543b743 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -224,7 +224,7 @@ lemma AffineIndependent.card_le_card_of_subset_affineSpan {s t : Finset V} have direction_le := AffineSubspace.direction_le (affineSpan_mono k hst) rw [AffineSubspace.affineSpan_coe, direction_affineSpan, direction_affineSpan, ← @Subtype.range_coe _ (s : Set V), ← @Subtype.range_coe _ (t : Set V)] at direction_le - have finrank_le := add_le_add_right (Submodule.finrank_le_finrank_of_le direction_le) 1 + have finrank_le := add_le_add_right (Submodule.finrank_mono direction_le) 1 -- We use `erw` to elide the difference between `↥s` and `↥(s : Set V)}` erw [hs.finrank_vectorSpan_add_one] at finrank_le simpa using finrank_le.trans <| finrank_vectorSpan_range_add_one_le _ _ @@ -371,7 +371,7 @@ alias ⟨Collinear.finrank_le_one, _⟩ := collinear_iff_finrank_le_one /-- A subset of a collinear set is collinear. -/ theorem Collinear.subset {s₁ s₂ : Set P} (hs : s₁ ⊆ s₂) (h : Collinear k s₂) : Collinear k s₁ := - (rank_le_of_submodule (vectorSpan k s₁) (vectorSpan k s₂) (vectorSpan_mono k hs)).trans h + (Submodule.rank_mono (vectorSpan_mono k hs)).trans h /-- The `vectorSpan` of collinear points is finite-dimensional. -/ theorem Collinear.finiteDimensional_vectorSpan {s : Set P} (h : Collinear k s) : @@ -634,7 +634,7 @@ alias ⟨Coplanar.finrank_le_two, _⟩ := coplanar_iff_finrank_le_two /-- A subset of a coplanar set is coplanar. -/ theorem Coplanar.subset {s₁ s₂ : Set P} (hs : s₁ ⊆ s₂) (h : Coplanar k s₂) : Coplanar k s₁ := - (rank_le_of_submodule (vectorSpan k s₁) (vectorSpan k s₂) (vectorSpan_mono k hs)).trans h + (Submodule.rank_mono (vectorSpan_mono k hs)).trans h /-- Collinear points are coplanar. -/ theorem Collinear.coplanar {s : Set P} (h : Collinear k s) : Coplanar k s := diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 9508cc5b7d60c..d3888ddec1094 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -285,11 +285,12 @@ theorem lift_rank_map_le (f : M →ₗ[R] M') (p : Submodule R M) : theorem rank_map_le (f : M →ₗ[R] M₁) (p : Submodule R M) : Module.rank R (p.map f) ≤ Module.rank R p := by simpa using lift_rank_map_le f p -theorem rank_le_of_submodule (s t : Submodule R M) (h : s ≤ t) : - Module.rank R s ≤ Module.rank R t := +lemma Submodule.rank_mono {s t : Submodule R M} (h : s ≤ t) : Module.rank R s ≤ Module.rank R t := (Submodule.inclusion h).rank_le_of_injective fun ⟨x, _⟩ ⟨y, _⟩ eq => Subtype.eq <| show x = y from Subtype.ext_iff_val.1 eq +@[deprecated (since := "2024-09-30")] alias rank_le_of_submodule := Submodule.rank_mono + /-- Two linearly equivalent vector spaces have the same dimension, a version with different universes. -/ theorem LinearEquiv.lift_rank_eq (f : M ≃ₗ[R] M') : @@ -331,9 +332,11 @@ theorem rank_range_of_surjective (f : M →ₗ[R] M') (h : Surjective f) : Module.rank R (LinearMap.range f) = Module.rank R M' := by rw [LinearMap.range_eq_top.2 h, rank_top] -theorem rank_submodule_le (s : Submodule R M) : Module.rank R s ≤ Module.rank R M := by +theorem Submodule.rank_le (s : Submodule R M) : Module.rank R s ≤ Module.rank R M := by rw [← rank_top R M] - exact rank_le_of_submodule _ _ le_top + exact rank_mono le_top + +@[deprecated (since := "2024-10-02")] alias rank_submodule_le := Submodule.rank_le theorem LinearMap.lift_rank_le_of_surjective (f : M →ₗ[R] M') (h : Surjective f) : lift.{v} (Module.rank R M') ≤ lift.{v'} (Module.rank R M) := by diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 2cf63df0ef942..2a3358623db3a 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -372,7 +372,7 @@ variable [StrongRankCondition R] /-- The dimension of a submodule is bounded by the dimension of the ambient space. -/ theorem Submodule.finrank_le [Module.Finite R M] (s : Submodule R M) : finrank R s ≤ finrank R M := - toNat_le_toNat (rank_submodule_le s) (rank_lt_aleph0 _ _) + toNat_le_toNat (Submodule.rank_le s) (rank_lt_aleph0 _ _) /-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ theorem Submodule.finrank_quotient_le [Module.Finite R M] (s : Submodule R M) : @@ -386,12 +386,12 @@ theorem Submodule.finrank_map_le finrank R (p.map f) ≤ finrank R p := finrank_le_finrank_of_rank_le_rank (lift_rank_map_le _ _) (rank_lt_aleph0 _ _) -theorem Submodule.finrank_le_finrank_of_le {s t : Submodule R M} [Module.Finite R t] (hst : s ≤ t) : +theorem Submodule.finrank_mono {s t : Submodule R M} [Module.Finite R t] (hst : s ≤ t) : finrank R s ≤ finrank R t := - calc - finrank R s = finrank R (s.comap t.subtype) := - (Submodule.comapSubtypeEquivOfLe hst).finrank_eq.symm - _ ≤ finrank R t := Submodule.finrank_le _ + Cardinal.toNat_le_toNat (Submodule.rank_mono hst) (rank_lt_aleph0 R ↥t) + +@[deprecated (since := "2024-09-30")] +alias Submodule.finrank_le_finrank_of_le := Submodule.finrank_mono end diff --git a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean index 418403e43f094..ca2cb869a8ec5 100644 --- a/Mathlib/LinearAlgebra/Dimension/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Dimension/LinearMap.lean @@ -34,7 +34,7 @@ abbrev rank (f : V →ₗ[K] V') : Cardinal := Module.rank K (LinearMap.range f) theorem rank_le_range (f : V →ₗ[K] V') : rank f ≤ Module.rank K V' := - rank_submodule_le _ + Submodule.rank_le _ theorem rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ Module.rank K V := rank_range_le _ @@ -46,7 +46,7 @@ theorem rank_zero [Nontrivial K] : rank (0 : V →ₗ[K] V') = 0 := by variable [AddCommGroup V''] [Module K V''] theorem rank_comp_le_left (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := by - refine rank_le_of_submodule _ _ ?_ + refine Submodule.rank_mono ?_ rw [LinearMap.range_comp] exact LinearMap.map_le_range @@ -82,7 +82,7 @@ variable [AddCommGroup V'] [Module K V'] theorem rank_add_le (f g : V →ₗ[K] V') : rank (f + g) ≤ rank f + rank g := calc rank (f + g) ≤ Module.rank K (LinearMap.range f ⊔ LinearMap.range g : Submodule K V') := by - refine rank_le_of_submodule _ _ ?_ + refine Submodule.rank_mono ?_ exact LinearMap.range_le_iff_comap.2 <| eq_top_iff'.2 fun x => show f x + g x ∈ (LinearMap.range f ⊔ LinearMap.range g : Submodule K V') from mem_sup.2 ⟨_, ⟨x, rfl⟩, _, ⟨x, rfl⟩, rfl⟩ diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 9b95709365e77..7802cfc81de05 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -148,7 +148,7 @@ instance finiteDimensional_submodule [FiniteDimensional K V] (S : Submodule K V) · exact iff_fg.1 (IsNoetherian.iff_rank_lt_aleph0.2 - (lt_of_le_of_lt (rank_submodule_le _) (_root_.rank_lt_aleph0 K V))) + ((Submodule.rank_le _).trans_lt (_root_.rank_lt_aleph0 K V))) · infer_instance /-- A quotient of a finite-dimensional space is also finite-dimensional. -/ @@ -374,8 +374,7 @@ theorem finiteDimensional_of_le {S₁ S₂ : Submodule K V} [FiniteDimensional K FiniteDimensional K S₁ := haveI : IsNoetherian K S₂ := iff_fg.2 inferInstance iff_fg.1 - (IsNoetherian.iff_rank_lt_aleph0.2 - (lt_of_le_of_lt (rank_le_of_submodule _ _ h) (rank_lt_aleph0 K S₂))) + (IsNoetherian.iff_rank_lt_aleph0.2 ((Submodule.rank_mono h).trans_lt (rank_lt_aleph0 K S₂))) /-- The inf of two submodules, the first finite-dimensional, is finite-dimensional. -/ @@ -694,21 +693,6 @@ noncomputable def fieldOfFiniteDimensional (F K : Type*) [Field F] [h : CommRing { divisionRingOfFiniteDimensional F K with toCommRing := h } end - -namespace Submodule - -section DivisionRing - -variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] - [Module K V₂] - -theorem finrank_mono [FiniteDimensional K V] : Monotone fun s : Submodule K V => finrank K s := - fun _ _ => finrank_le_finrank_of_le - -end DivisionRing - -end Submodule - section DivisionRing variable [DivisionRing K] [AddCommGroup V] [Module K V] diff --git a/Mathlib/LinearAlgebra/LinearDisjoint.lean b/Mathlib/LinearAlgebra/LinearDisjoint.lean index 3eb097c1e65c0..f290e4e777d86 100644 --- a/Mathlib/LinearAlgebra/LinearDisjoint.lean +++ b/Mathlib/LinearAlgebra/LinearDisjoint.lean @@ -512,7 +512,7 @@ if any two elements of `↥(M ⊓ N)` are commutative, then the rank of `↥(M theorem rank_inf_le_one_of_commute_of_flat (hf : Module.Flat R M ∨ Module.Flat R N) (hc : ∀ (m n : ↥(M ⊓ N)), Commute m.1 n.1) : Module.rank R ↥(M ⊓ N) ≤ 1 := by nontriviality R - refine rank_le fun s h ↦ ?_ + refine _root_.rank_le fun s h ↦ ?_ by_contra hs rw [not_le, ← Fintype.card_coe, Fintype.one_lt_card_iff_nontrivial] at hs obtain ⟨a, b, hab⟩ := hs.exists_pair_ne From f84123514f85483265d95458b3203f81651de3f2 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 3 Oct 2024 05:41:53 +0000 Subject: [PATCH 143/174] chore(Algebra/BigOperators/Module): restore `conv` (#17356) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/BigOperators/Module.lean | 19 ++----------------- 1 file changed, 2 insertions(+), 17 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Module.lean b/Mathlib/Algebra/BigOperators/Module.lean index fd9ddbeeb76ee..2329df72948f4 100644 --- a/Mathlib/Algebra/BigOperators/Module.lean +++ b/Mathlib/Algebra/BigOperators/Module.lean @@ -31,24 +31,9 @@ theorem sum_Ico_by_parts (hmn : m < n) : rw [← sum_Ico_sub_bot _ hmn, ← sum_Ico_succ_sub_top _ (Nat.le_sub_one_of_lt hmn), Nat.sub_add_cancel (pos_of_gt hmn), sub_add_cancel] rw [sum_eq_sum_Ico_succ_bot hmn] - -- Porting note: the following used to be done with `conv` - have h₃ : (Finset.sum (Ico (m + 1) n) fun i => f i • g i) = - (Finset.sum (Ico (m + 1) n) fun i => - f i • ((Finset.sum (Finset.range (i + 1)) g) - - (Finset.sum (Finset.range i) g))) := by - congr; funext; rw [← sum_range_succ_sub_sum g] - rw [h₃] + conv in (occs := 3) (f _ • g _) => rw [← sum_range_succ_sub_sum g] simp_rw [smul_sub, sum_sub_distrib, h₂, h₁] - -- Porting note: the following used to be done with `conv` - have h₄ : ((((Finset.sum (Ico m (n - 1)) fun i => f i • Finset.sum (range (i + 1)) fun i => g i) + - f (n - 1) • Finset.sum (range n) fun i => g i) - - f m • Finset.sum (range (m + 1)) fun i => g i) - - Finset.sum (Ico m (n - 1)) fun i => f (i + 1) • Finset.sum (range (i + 1)) fun i => g i) = - f (n - 1) • (range n).sum g - f m • (range (m + 1)).sum g + - Finset.sum (Ico m (n - 1)) (fun i => f i • (range (i + 1)).sum g - - f (i + 1) • (range (i + 1)).sum g) := by - rw [← add_sub, add_comm, ← add_sub, ← sum_sub_distrib] - rw [h₄] + conv_lhs => congr; rfl; rw [← add_sub, add_comm, ← add_sub, ← sum_sub_distrib] have : ∀ i, f i • G (i + 1) - f (i + 1) • G (i + 1) = -((f (i + 1) - f i) • G (i + 1)) := by intro i rw [sub_smul] From 39e0665cd3f129042677a5dab6081f0ada9b9d2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Thu, 3 Oct 2024 07:38:17 +0000 Subject: [PATCH 144/174] feat(FieldTheory): add lemmas about minpoly (#15723) Co-authored-by: Filippo A. E. Nuccio --- Mathlib/FieldTheory/Adjoin.lean | 9 +++++++++ .../Minpoly/IsIntegrallyClosed.lean | 20 +++++++++++++++++++ 2 files changed, 29 insertions(+) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 07881787d2645..e4701e89934ea 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -1114,6 +1114,15 @@ theorem _root_.minpoly.degree_le (x : L) [FiniteDimensional K L] : (minpoly K x).degree ≤ finrank K L := degree_le_of_natDegree_le (minpoly.natDegree_le x) +/-- If `x : L` is an integral element in a field extension `L` over `K`, then the degree of the + minimal polynomial of `x` over `K` divides `[L : K]`.-/ +theorem _root_.minpoly.degree_dvd {x : L} (hx : IsIntegral K x) : + (minpoly K x).natDegree ∣ FiniteDimensional.finrank K L := by + rw [dvd_iff_exists_eq_mul_left, ← IntermediateField.adjoin.finrank hx] + use FiniteDimensional.finrank K⟮x⟯ L + rw [eq_comm, mul_comm] + exact FiniteDimensional.finrank_mul_finrank _ _ _ + -- TODO: generalize to `Sort` /-- A compositum of algebraic extensions is algebraic -/ theorem isAlgebraic_iSup {ι : Type*} {t : ι → IntermediateField K L} diff --git a/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean b/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean index 5abd2d46ef12c..3db3147d29b05 100644 --- a/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean +++ b/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean @@ -186,4 +186,24 @@ theorem _root_.PowerBasis.ofGenMemAdjoin'_gen (B : PowerBasis R S) (hint : IsInt end AdjoinRoot +section Subring + +variable {K L : Type*} [Field K] [Field L] [Algebra K L] + +variable (A : Subring K) [IsIntegrallyClosed A] [IsFractionRing A K] + +-- Implementation note: `inferInstance` does not work for these. +instance : Algebra A (integralClosure A L) := Subalgebra.algebra (integralClosure A L) +instance : SMul A (integralClosure A L) := Algebra.toSMul +instance : IsScalarTower A ((integralClosure A L)) L := + IsScalarTower.subalgebra' A L L (integralClosure A L) + +/-- The minimal polynomial of `x : L` over `K` agrees with its minimal polynomial over the +integrally closed subring `A`. -/ +theorem ofSubring (x : integralClosure A L) : + Polynomial.map (algebraMap A K) (minpoly A x) = minpoly K (x : L) := + eq_comm.mpr (isIntegrallyClosed_eq_field_fractions K L (IsIntegralClosure.isIntegral A L x)) + +end Subring + end minpoly From 23761ebb9bdc2b099b842165fe6f804f69ceeac1 Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Thu, 3 Oct 2024 10:05:46 +0000 Subject: [PATCH 145/174] refactor(SetTheory/Ordinal/Basic): deprecate Ordinal.omega in favor of Ordinal.omega0 (#17158) See [zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.CF.89_.20.CE.B1.20function) for the discussion. --- Mathlib/SetTheory/Cardinal/Cofinality.lean | 9 +- Mathlib/SetTheory/Cardinal/Ordinal.lean | 17 ++- Mathlib/SetTheory/Game/Nim.lean | 4 +- Mathlib/SetTheory/Game/Short.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 103 +++++++++----- Mathlib/SetTheory/Ordinal/Basic.lean | 28 ++-- Mathlib/SetTheory/Ordinal/Exponential.lean | 4 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 108 +++++++++------ Mathlib/SetTheory/Ordinal/Notation.lean | 132 ++++++++++-------- Mathlib/SetTheory/Ordinal/Principal.lean | 154 +++++++++++++-------- scripts/no_lints_prime_decls.txt | 4 +- 11 files changed, 350 insertions(+), 215 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 035f8671aa50d..53bfb76a4a06e 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -680,11 +680,14 @@ theorem aleph_cof {o : Ordinal} (ho : o.IsLimit) : (aleph o).ord.cof = o.cof := aleph_isNormal.cof_eq ho @[simp] -theorem cof_omega : cof ω = ℵ₀ := - (aleph0_le_cof.2 omega_isLimit).antisymm' <| by - rw [← card_omega] +theorem cof_omega0 : cof ω = ℵ₀ := + (aleph0_le_cof.2 omega0_isLimit).antisymm' <| by + rw [← card_omega0] apply cof_le_card +@[deprecated (since := "2024-09-30")] +alias cof_omega := cof_omega0 + theorem cof_eq' (r : α → α → Prop) [IsWellOrder α r] (h : IsLimit (type r)) : ∃ S : Set α, (∀ a, ∃ b ∈ S, r a b) ∧ #S = cof (type r) := let ⟨S, H, e⟩ := cof_eq r diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index f605981cb3c6b..fe8a2ac903837 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -61,7 +61,7 @@ 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 omega_isLimit + exact omega0_isLimit theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := toType_noMax_of_succ_lt (ord_isLimit h).2 @@ -217,11 +217,14 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, exact fun a ha => le_ciSup (bddAbove_of_small _) (⟨a, ha⟩ : Iio o) @[simp] -theorem aleph'_omega : aleph' ω = ℵ₀ := +theorem aleph'_omega0 : aleph' ω = ℵ₀ := eq_of_forall_ge_iff fun c => by - simp only [aleph'_le_of_limit omega_isLimit, lt_omega, exists_imp, aleph0_le] + simp only [aleph'_le_of_limit omega0_isLimit, 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 + set_option linter.deprecated false in /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @[deprecated aleph' (since := "2024-08-28")] @@ -256,7 +259,7 @@ theorem aleph_succ (o : Ordinal) : aleph (succ o) = succ (aleph o) := by rw [aleph_eq_aleph', add_succ, aleph'_succ, aleph_eq_aleph'] @[simp] -theorem aleph_zero : aleph 0 = ℵ₀ := by rw [aleph_eq_aleph', add_zero, aleph'_omega] +theorem aleph_zero : aleph 0 = ℵ₀ := by rw [aleph_eq_aleph', add_zero, aleph'_omega0] theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : aleph o = ⨆ a : Iio o, aleph a := by apply le_antisymm _ (ciSup_le' _) @@ -264,13 +267,13 @@ theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : aleph o = ⨆ a : Iio o, al refine ciSup_mono' (bddAbove_of_small _) ?_ rintro ⟨i, hi⟩ cases' lt_or_le i ω with h h - · rcases lt_omega.1 h with ⟨n, rfl⟩ + · 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'_omega, aleph'_le] +theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by rw [← aleph'_omega0, aleph'_le] theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ aleph o := by rw [aleph_eq_aleph', aleph0_le_aleph'] @@ -1442,7 +1445,7 @@ scoped notation "ω_" o => ord <| aleph o -/ scoped notation "ω₁" => ord <| aleph 1 -lemma omega_lt_omega1 : ω < ω₁ := ord_aleph0.symm.trans_lt (ord_lt_ord.mpr (aleph0_lt_aleph_one)) +lemma omega0_lt_omega1 : ω < ω₁ := ord_aleph0.symm.trans_lt (ord_lt_ord.mpr (aleph0_lt_aleph_one)) section OrdinalIndices /-! diff --git a/Mathlib/SetTheory/Game/Nim.lean b/Mathlib/SetTheory/Game/Nim.lean index dea7c7401922c..46bd9ae9c794e 100644 --- a/Mathlib/SetTheory/Game/Nim.lean +++ b/Mathlib/SetTheory/Game/Nim.lean @@ -343,7 +343,7 @@ theorem grundyValue_nim_add_nim (n m : ℕ) : grundyValue (nim.{u} n + nim.{u} m all_goals intro j have hj := toLeftMovesNim_symm_lt j - obtain ⟨k, hk⟩ := lt_omega.1 (hj.trans (nat_lt_omega _)) + obtain ⟨k, hk⟩ := lt_omega0.1 (hj.trans (nat_lt_omega0 _)) rw [hk, Nat.cast_lt] at hj have := hj.ne have := hj -- The termination checker doesn't work without this. @@ -354,7 +354,7 @@ theorem grundyValue_nim_add_nim (n m : ℕ) : grundyValue (nim.{u} n + nim.{u} m -- For any `k < n ^^^ m`, either `nim (k ^^^ m) + nim m` or `nim n + nim (k ^^^ n)` is a left -- option with Grundy value `k`. · intro k hk - obtain ⟨k, rfl⟩ := Ordinal.lt_omega.1 (hk.trans (Ordinal.nat_lt_omega _)) + obtain ⟨k, rfl⟩ := Ordinal.lt_omega0.1 (hk.trans (Ordinal.nat_lt_omega0 _)) rw [Set.mem_Iio, Nat.cast_lt] at hk obtain hk | hk := Nat.lt_xor_cases hk <;> rw [← natCast_lt] at hk · use toLeftMovesAdd (Sum.inl (toLeftMovesNim ⟨_, hk⟩)) diff --git a/Mathlib/SetTheory/Game/Short.lean b/Mathlib/SetTheory/Game/Short.lean index 1d55f6247f452..03c8ecbda2c5b 100644 --- a/Mathlib/SetTheory/Game/Short.lean +++ b/Mathlib/SetTheory/Game/Short.lean @@ -136,7 +136,7 @@ def moveRightShort' {xl xr} (xL xR) [S : Short (mk xl xr xL xR)] (j : xr) : Shor attribute [local instance] moveRightShort' -theorem short_birthday (x : PGame.{u}) : [Short x] → x.birthday < Ordinal.omega := by +theorem short_birthday (x : PGame.{u}) : [Short x] → x.birthday < Ordinal.omega0 := by -- Porting note: Again `induction` is used instead of `pgame_wf_tac` induction x with | mk xl xr xL xR ihl ihr => diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index c53d3f62b99c5..7c32c24c42d88 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -539,9 +539,9 @@ theorem sub_isLimit {a b} (l : IsLimit a) (h : b < a) : IsLimit (a - b) := rw [lt_sub, add_succ]; exact l.2 _ (lt_sub.1 h)⟩ -- @[simp] -- Porting note (#10618): simp can prove this -theorem one_add_omega : 1 + ω = ω := by +theorem one_add_omega0 : 1 + ω = ω := by refine le_antisymm ?_ (le_add_left _ _) - rw [omega, ← lift_one.{0}, ← lift_add, lift_le, ← type_unit, ← type_sum_lex] + rw [omega0, ← lift_one.{0}, ← lift_add, lift_le, ← type_unit, ← type_sum_lex] refine ⟨RelEmbedding.collapse (RelEmbedding.ofMonotone ?_ ?_)⟩ · apply Sum.rec · exact fun _ => 0 @@ -550,9 +550,15 @@ theorem one_add_omega : 1 + ω = ω := by cases a <;> cases b <;> intro H <;> cases' H with _ _ H _ _ H <;> [exact H.elim; exact Nat.succ_pos _; exact Nat.succ_lt_succ H] +@[deprecated (since := "2024-09-30")] +alias one_add_omega := one_add_omega0 + @[simp] -theorem one_add_of_omega_le {o} (h : ω ≤ o) : 1 + o = o := by - rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, one_add_omega] +theorem one_add_of_omega0_le {o} (h : ω ≤ o) : 1 + o = o := by + rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, one_add_omega0] + +@[deprecated (since := "2024-09-30")] +alias one_add_of_omega_le := one_add_of_omega0_le /-! ### Multiplication of ordinals -/ @@ -2252,7 +2258,7 @@ theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : end Ordinal -/-! ### Properties of `omega` -/ +/-! ### Properties of ω -/ namespace Cardinal @@ -2269,7 +2275,7 @@ theorem ord_aleph0 : ord.{u} ℵ₀ = ω := @[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_omega_le] + rw [add_comm, ← card_ord c, ← card_one, ← card_add, one_add_of_omega0_le] rwa [← ord_aleph0, ord_le_ord] end Cardinal @@ -2281,34 +2287,56 @@ theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : -- Porting note: `bex_def` is required. rw [← IsNormal.bsup_eq.{u, u} (add_isNormal b) h, lt_bsup, bex_def] -theorem lt_omega {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by +theorem lt_omega0 {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by simp_rw [← Cardinal.ord_aleph0, Cardinal.lt_ord, lt_aleph0, card_eq_nat] -theorem nat_lt_omega (n : ℕ) : ↑n < ω := - lt_omega.2 ⟨_, rfl⟩ +@[deprecated (since := "2024-09-30")] +alias lt_omega := lt_omega0 + +theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := + lt_omega0.2 ⟨_, rfl⟩ +@[deprecated (since := "2024-09-30")] +alias nat_lt_omega := nat_lt_omega0 + +theorem omega0_pos : 0 < ω := + nat_lt_omega0 0 + +@[deprecated (since := "2024-09-30")] theorem omega_pos : 0 < ω := - nat_lt_omega 0 + nat_lt_omega0 0 + +theorem omega0_ne_zero : ω ≠ 0 := + omega0_pos.ne' -theorem omega_ne_zero : ω ≠ 0 := - omega_pos.ne' +@[deprecated (since := "2024-09-30")] +alias omega_ne_zero := omega0_ne_zero -theorem one_lt_omega : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega 1 +theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega0 1 -theorem omega_isLimit : IsLimit ω := - ⟨omega_ne_zero, fun o h => by - let ⟨n, e⟩ := lt_omega.1 h - rw [e]; exact nat_lt_omega (n + 1)⟩ +@[deprecated (since := "2024-09-30")] +alias one_lt_omega := one_lt_omega0 -theorem omega_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := - ⟨fun h n => (nat_lt_omega _).le.trans h, fun H => +theorem omega0_isLimit : 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-09-30")] +alias omega_isLimit := omega0_isLimit + +theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := + ⟨fun h n => (nat_lt_omega0 _).le.trans h, fun H => le_of_forall_lt fun a h => by - let ⟨n, e⟩ := lt_omega.1 h + let ⟨n, e⟩ := lt_omega0.1 h rw [e, ← succ_le_iff]; exact H (n + 1)⟩ +@[deprecated (since := "2024-09-30")] +alias omega_le := omega0_le + @[simp] theorem iSup_natCast : iSup Nat.cast = ω := - (Ordinal.iSup_le fun n => (nat_lt_omega n).le).antisymm <| omega_le.2 <| Ordinal.le_iSup _ + (Ordinal.iSup_le fun n => (nat_lt_omega0 n).le).antisymm <| omega0_le.2 <| Ordinal.le_iSup _ set_option linter.deprecated false in @[deprecated iSup_natCast (since := "2024-04-17")] @@ -2322,24 +2350,30 @@ theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o | 0 => lt_of_le_of_ne (Ordinal.zero_le o) h.1.symm | n + 1 => h.2 _ (nat_lt_limit h n) -theorem omega_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := - omega_le.2 fun n => le_of_lt <| nat_lt_limit h n +theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := + omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n -theorem isLimit_iff_omega_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by +@[deprecated (since := "2024-09-30")] +alias omega_le_of_isLimit := omega0_le_of_isLimit + +theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by 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 omega_ne_zero, ← succ_le_iff, le_div omega_ne_zero, mul_succ, - add_le_of_limit omega_isLimit] + rw [← div_lt omega0_ne_zero, ← succ_le_iff, le_div omega0_ne_zero, mul_succ, + add_le_of_limit omega0_isLimit] intro b hb - rcases lt_omega.1 hb with ⟨n, rfl⟩ + rcases lt_omega0.1 hb with ⟨n, rfl⟩ exact (add_le_add_right (mul_div_le _ _) _).trans (lt_sub.1 <| nat_lt_limit (sub_isLimit l hx) _).le · rcases h with ⟨a0, b, rfl⟩ - refine mul_isLimit_left omega_isLimit (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0) + refine mul_isLimit_left omega0_isLimit (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0) intro e simp only [e, mul_zero] +@[deprecated (since := "2024-09-30")] +alias isLimit_iff_omega_dvd := isLimit_iff_omega0_dvd + theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) (IH : ∀ c' < c, (a + b) * succ c' = a * succ c' + b) : (a + b) * c = a * c := le_antisymm @@ -2378,24 +2412,27 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a by_contra! hb exact (h _ hb).ne H -theorem IsNormal.apply_omega {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : +theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] +@[deprecated (since := "2024-09-30")] +alias IsNormal.apply_omega := IsNormal.apply_omega0 + @[simp] theorem iSup_add_nat (o : Ordinal) : ⨆ n : ℕ, o + n = o + ω := - (add_isNormal o).apply_omega + (add_isNormal o).apply_omega0 set_option linter.deprecated false in @[deprecated iSup_add_nat (since := "2024-08-27")] theorem sup_add_nat (o : Ordinal) : (sup fun n : ℕ => o + n) = o + ω := - (add_isNormal o).apply_omega + (add_isNormal o).apply_omega0 @[simp] theorem iSup_mul_nat (o : Ordinal) : ⨆ n : ℕ, o * n = o * ω := by rcases eq_zero_or_pos o with (rfl | ho) · rw [zero_mul] exact iSup_eq_zero_iff.2 fun n => zero_mul (n : Ordinal) - · exact (mul_isNormal ho).apply_omega + · exact (mul_isNormal ho).apply_omega0 set_option linter.deprecated false in @[deprecated iSup_add_nat (since := "2024-08-27")] @@ -2403,7 +2440,7 @@ theorem sup_mul_nat (o : Ordinal) : (sup fun n : ℕ => o * n) = o * ω := by rcases eq_zero_or_pos o with (rfl | ho) · rw [zero_mul] exact sup_eq_zero_iff.2 fun n => zero_mul (n : Ordinal) - · exact (mul_isNormal ho).apply_omega + · exact (mul_isNormal ho).apply_omega0 end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index d9a0cb4a39c91..895d453ab84d9 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -31,9 +31,10 @@ initial segment (or, equivalently, in any way). This total order is well founded `Ordinal.liftInitialSeg`. For a version registering that it is a principal segment embedding if `u < v`, see `Ordinal.liftPrincipalSeg`. -* `Ordinal.omega` or `ω` is the order type of `ℕ`. This definition is universe polymorphic: - `Ordinal.omega.{u} : Ordinal.{u}` (contrast with `ℕ : Type`, which lives in a specific - universe). In some cases the universe level has to be given explicitly. +* `Ordinal.omega0` or `ω` is the order type of `ℕ`. It is called this to match `Cardinal.aleph0` + and so that the omega function can be named `Ordinal.omega`. This definition is universe + polymorphic: `Ordinal.omega0.{u} : Ordinal.{u}` (contrast with `ℕ : Type`, which lives in + a specific universe). In some cases the universe level has to be given explicitly. * `o₁ + o₂` is the order on the disjoint union of `o₁` and `o₂` obtained by declaring that every element of `o₁` is smaller than every element of `o₂`. @@ -701,29 +702,38 @@ set_option linter.deprecated false in theorem lift.initialSeg_coe : (lift.initialSeg.{u, v} : Ordinal → Ordinal) = lift.{v, u} := rfl -/-! ### The first infinite ordinal `omega` -/ +/-! ### The first infinite ordinal ω -/ /-- `ω` is the first infinite ordinal, defined as the order type of `ℕ`. -/ -def omega : Ordinal.{u} := +def omega0 : Ordinal.{u} := lift <| @type ℕ (· < ·) _ +@[deprecated Ordinal.omega0 (since := "2024-09-26")] +alias omega := omega0 + @[inherit_doc] -scoped notation "ω" => Ordinal.omega +scoped notation "ω" => Ordinal.omega0 -/-- Note that the presence of this lemma makes `simp [omega]` form a loop. -/ +/-- Note that the presence of this lemma makes `simp [omega0]` form a loop. -/ @[simp] theorem type_nat_lt : @type ℕ (· < ·) _ = ω := (lift_id _).symm @[simp] -theorem card_omega : card ω = ℵ₀ := +theorem card_omega0 : card ω = ℵ₀ := rfl +@[deprecated (since := "2024-09-30")] +alias card_omega := card_omega0 + @[simp] -theorem lift_omega : lift ω = ω := +theorem lift_omega0 : lift ω = ω := lift_lift _ +@[deprecated (since := "2024-09-30")] +alias lift_omega := lift_omega0 + /-! ### Definition and first properties of addition on ordinals diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index d134e192623ac..f5d3ed03ea35f 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -403,7 +403,7 @@ theorem natCast_opow (m : ℕ) : ∀ n : ℕ, ↑(m ^ n : ℕ) = (m : Ordinal) ^ theorem iSup_pow {o : Ordinal} (ho : 0 < o) : ⨆ n : ℕ, o ^ n = o ^ ω := by simp_rw [← opow_natCast] rcases (one_le_iff_pos.2 ho).lt_or_eq with ho₁ | rfl - · exact (opow_isNormal ho₁).apply_omega + · exact (opow_isNormal ho₁).apply_omega0 · rw [one_opow] refine le_antisymm (Ordinal.iSup_le fun n => by rw [one_opow]) ?_ convert Ordinal.le_iSup _ 0 @@ -414,7 +414,7 @@ set_option linter.deprecated false in theorem sup_opow_nat {o : Ordinal} (ho : 0 < o) : (sup fun n : ℕ => o ^ n) = o ^ ω := by simp_rw [← opow_natCast] rcases (one_le_iff_pos.2 ho).lt_or_eq with ho₁ | rfl - · exact (opow_isNormal ho₁).apply_omega + · exact (opow_isNormal ho₁).apply_omega0 · rw [one_opow] refine le_antisymm (sup_le fun n => by rw [one_opow]) ?_ convert le_sup (fun n : ℕ => 1 ^ (n : Ordinal)) 0 diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 235483b8420b9..4fe7f81c09218 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -21,8 +21,8 @@ Moreover, we prove some lemmas about the fixed points of specific normal functio * `nfpFamily`, `nfpBFamily`, `nfp`: the next fixed point of a (family of) normal function(s). * `fp_family_unbounded`, `fp_bfamily_unbounded`, `fp_unbounded`: the (common) fixed points of a (family of) normal function(s) are unbounded in the ordinals. -* `deriv_add_eq_mul_omega_add`: a characterization of the derivative of addition. -* `deriv_mul_eq_opow_omega_mul`: a characterization of the derivative of multiplication. +* `deriv_add_eq_mul_omega0_add`: a characterization of the derivative of addition. +* `deriv_mul_eq_opow_omega0_mul`: a characterization of the derivative of multiplication. -/ @@ -521,20 +521,23 @@ end /-! ### Fixed points of addition -/ @[simp] -theorem nfp_add_zero (a) : nfp (a + ·) 0 = a * omega := by +theorem nfp_add_zero (a) : nfp (a + ·) 0 = a * ω := by simp_rw [← iSup_iterate_eq_nfp, ← iSup_mul_nat] congr; funext n induction' n with n hn · rw [Nat.cast_zero, mul_zero, iterate_zero_apply] · rw [iterate_succ_apply', Nat.add_comm, Nat.cast_add, Nat.cast_one, mul_one_add, hn] -theorem nfp_add_eq_mul_omega {a b} (hba : b ≤ a * omega) : nfp (a + ·) b = a * omega := by +theorem nfp_add_eq_mul_omega0 {a b} (hba : b ≤ a * ω) : nfp (a + ·) b = a * ω := by apply le_antisymm (nfp_le_fp (add_isNormal a).monotone hba _) · rw [← nfp_add_zero] exact nfp_monotone (add_isNormal a).monotone (Ordinal.zero_le b) - · dsimp; rw [← mul_one_add, one_add_omega] + · dsimp; rw [← mul_one_add, one_add_omega0] -theorem add_eq_right_iff_mul_omega_le {a b : Ordinal} : a + b = b ↔ a * omega ≤ b := by +@[deprecated (since := "2024-09-30")] +alias nfp_add_eq_mul_omega := nfp_add_eq_mul_omega0 + +theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω ≤ b := by refine ⟨fun h => ?_, fun h => ?_⟩ · rw [← nfp_add_zero a, ← deriv_zero_right] cases' (add_isNormal a).fp_iff_deriv.1 h with c hc @@ -542,25 +545,34 @@ theorem add_eq_right_iff_mul_omega_le {a b : Ordinal} : a + b = b ↔ a * omega exact (deriv_isNormal _).monotone (Ordinal.zero_le _) · have := Ordinal.add_sub_cancel_of_le h nth_rw 1 [← this] - rwa [← add_assoc, ← mul_one_add, one_add_omega] + rwa [← add_assoc, ← mul_one_add, one_add_omega0] + +@[deprecated (since := "2024-09-30")] +alias add_eq_right_iff_mul_omega_le := add_eq_right_iff_mul_omega0_le -theorem add_le_right_iff_mul_omega_le {a b : Ordinal} : a + b ≤ b ↔ a * omega ≤ b := by - rw [← add_eq_right_iff_mul_omega_le] +theorem add_le_right_iff_mul_omega0_le {a b : Ordinal} : a + b ≤ b ↔ a * ω ≤ b := by + rw [← add_eq_right_iff_mul_omega0_le] exact (add_isNormal a).le_iff_eq -theorem deriv_add_eq_mul_omega_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * omega + b := by +@[deprecated (since := "2024-09-30")] +alias add_le_right_iff_mul_omega_le := add_le_right_iff_mul_omega0_le + +theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * ω + b := by revert b rw [← funext_iff, IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) (add_isNormal _)] refine ⟨?_, fun a h => ?_⟩ · rw [deriv_zero_right, add_zero] exact nfp_add_zero a · rw [deriv_succ, h, add_succ] - exact nfp_eq_self (add_eq_right_iff_mul_omega_le.2 ((le_add_right _ _).trans (le_succ _))) + exact nfp_eq_self (add_eq_right_iff_mul_omega0_le.2 ((le_add_right _ _).trans (le_succ _))) + +@[deprecated (since := "2024-09-30")] +alias deriv_add_eq_mul_omega_add := deriv_add_eq_mul_omega0_add /-! ### Fixed points of multiplication -/ @[simp] -theorem nfp_mul_one {a : Ordinal} (ha : 0 < a) : nfp (a * ·) 1 = (a^omega) := by +theorem nfp_mul_one {a : Ordinal} (ha : 0 < a) : nfp (a * ·) 1 = (a ^ ω) := by rw [← iSup_iterate_eq_nfp, ← iSup_pow ha] congr funext n @@ -575,22 +587,25 @@ theorem nfp_mul_zero (a : Ordinal) : nfp (a * ·) 0 = 0 := by induction' n with n hn; · rfl dsimp only; rwa [iterate_succ_apply, mul_zero] -theorem nfp_mul_eq_opow_omega {a b : Ordinal} (hb : 0 < b) (hba : b ≤ (a^omega)) : - nfp (a * ·) b = (a^omega.{u}) := by +theorem nfp_mul_eq_opow_omega0 {a b : Ordinal} (hb : 0 < b) (hba : b ≤ (a ^ ω)) : + nfp (a * ·) b = (a ^ (ω : Ordinal.{u})) := by rcases eq_zero_or_pos a with ha | ha - · rw [ha, zero_opow omega_ne_zero] at hba ⊢ + · rw [ha, zero_opow omega0_ne_zero] at hba ⊢ simp_rw [Ordinal.le_zero.1 hba, zero_mul] exact nfp_zero_left 0 apply le_antisymm · apply nfp_le_fp (mul_isNormal ha).monotone hba - rw [← opow_one_add, one_add_omega] + rw [← opow_one_add, one_add_omega0] rw [← nfp_mul_one ha] exact nfp_monotone (mul_isNormal ha).monotone (one_le_iff_pos.2 hb) -theorem eq_zero_or_opow_omega_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = b) : - b = 0 ∨ (a^omega.{u}) ≤ b := by +@[deprecated (since := "2024-09-30")] +alias nfp_mul_eq_opow_omega := nfp_mul_eq_opow_omega0 + +theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = b) : + b = 0 ∨ (a ^ (ω : Ordinal.{u})) ≤ b := by rcases eq_zero_or_pos a with ha | ha - · rw [ha, zero_opow omega_ne_zero] + · rw [ha, zero_opow omega0_ne_zero] exact Or.inr (Ordinal.zero_le b) rw [or_iff_not_imp_left] intro hb @@ -598,51 +613,66 @@ theorem eq_zero_or_opow_omega_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = rw [← Ne, ← one_le_iff_ne_zero] at hb exact nfp_le_fp (mul_isNormal ha).monotone hb (le_of_eq hab) -theorem mul_eq_right_iff_opow_omega_dvd {a b : Ordinal} : a * b = b ↔ (a^omega) ∣ b := by +@[deprecated (since := "2024-09-30")] +alias eq_zero_or_opow_omega_le_of_mul_eq_right := eq_zero_or_opow_omega0_le_of_mul_eq_right + +theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ (a ^ ω) ∣ b := by rcases eq_zero_or_pos a with ha | ha - · rw [ha, zero_mul, zero_opow omega_ne_zero, zero_dvd_iff] + · rw [ha, zero_mul, zero_opow omega0_ne_zero, zero_dvd_iff] exact eq_comm refine ⟨fun hab => ?_, fun h => ?_⟩ · rw [dvd_iff_mod_eq_zero] - rw [← div_add_mod b (a^omega), mul_add, ← mul_assoc, ← opow_one_add, one_add_omega, + rw [← div_add_mod b (a ^ ω), mul_add, ← mul_assoc, ← opow_one_add, one_add_omega0, add_left_cancel] at hab - cases' eq_zero_or_opow_omega_le_of_mul_eq_right hab with hab hab + cases' eq_zero_or_opow_omega0_le_of_mul_eq_right hab with hab hab · exact hab - refine (not_lt_of_le hab (mod_lt b (opow_ne_zero omega ?_))).elim + refine (not_lt_of_le hab (mod_lt b (opow_ne_zero ω ?_))).elim rwa [← Ordinal.pos_iff_ne_zero] cases' h with c hc - rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega] + rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega0] + +@[deprecated (since := "2024-09-30")] +alias mul_eq_right_iff_opow_omega_dvd := mul_eq_right_iff_opow_omega0_dvd -theorem mul_le_right_iff_opow_omega_dvd {a b : Ordinal} (ha : 0 < a) : - a * b ≤ b ↔ (a^omega) ∣ b := by - rw [← mul_eq_right_iff_opow_omega_dvd] +theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : + a * b ≤ b ↔ (a ^ ω) ∣ b := by + rw [← mul_eq_right_iff_opow_omega0_dvd] exact (mul_isNormal ha).le_iff_eq -theorem nfp_mul_opow_omega_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) (hca : c ≤ (a^omega)) : - nfp (a * ·) ((a^omega) * b + c) = (a^omega.{u}) * succ b := by +@[deprecated (since := "2024-09-30")] +alias mul_le_right_iff_opow_omega_dvd := mul_le_right_iff_opow_omega0_dvd + +theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) + (hca : c ≤ a ^ ω) : nfp (a * ·) (a ^ ω * b + c) = (a ^ (ω : Ordinal.{u})) * succ b := by apply le_antisymm · apply nfp_le_fp (mul_isNormal ha).monotone · rw [mul_succ] apply add_le_add_left hca - · dsimp only; rw [← mul_assoc, ← opow_one_add, one_add_omega] - · cases' mul_eq_right_iff_opow_omega_dvd.1 ((mul_isNormal ha).nfp_fp ((a^omega) * b + c)) with + · dsimp only; rw [← mul_assoc, ← opow_one_add, one_add_omega0] + · cases' mul_eq_right_iff_opow_omega0_dvd.1 ((mul_isNormal ha).nfp_fp ((a ^ ω) * b + c)) with d hd rw [hd] apply mul_le_mul_left' - have := le_nfp (Mul.mul a) ((a^omega) * b + c) + have := le_nfp (a * ·) (a ^ ω * b + c) erw [hd] at this - have := (add_lt_add_left hc ((a^omega) * b)).trans_le this - rw [add_zero, mul_lt_mul_iff_left (opow_pos omega ha)] at this + have := (add_lt_add_left hc (a ^ ω * b)).trans_le this + rw [add_zero, mul_lt_mul_iff_left (opow_pos ω ha)] at this rwa [succ_le_iff] -theorem deriv_mul_eq_opow_omega_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : - deriv (a * ·) b = (a^omega) * b := by +@[deprecated (since := "2024-09-30")] +alias nfp_mul_opow_omega_add := nfp_mul_opow_omega0_add + +theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : + deriv (a * ·) b = (a ^ ω) * b := by revert b rw [← funext_iff, - IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) (mul_isNormal (opow_pos omega ha))] + IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) (mul_isNormal (opow_pos ω ha))] refine ⟨?_, fun c h => ?_⟩ · dsimp only; rw [deriv_zero_right, nfp_mul_zero, mul_zero] · rw [deriv_succ, h] - exact nfp_mul_opow_omega_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) + exact nfp_mul_opow_omega0_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) + +@[deprecated (since := "2024-09-30")] +alias deriv_mul_eq_opow_omega_mul := deriv_mul_eq_opow_omega0_mul end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index d32c08e3c86d1..fd54a141e2131 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -139,12 +139,15 @@ 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 omega_le_oadd (e n a) : ω ^ repr e ≤ repr (oadd e n a) := by +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) omega_pos).2 (natCast_le.2 n.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e) omega0_pos).2 (natCast_le.2 n.2) + +@[deprecated (since := "2024-09-30")] +alias omega_le_oadd := omega0_le_oadd theorem oadd_pos (e n a) : 0 < oadd e n a := - @lt_of_lt_of_le _ _ _ (ω ^ repr e) _ (opow_pos (repr e) omega_pos) (omega_le_oadd e n a) + @lt_of_lt_of_le _ _ _ (ω ^ repr e) _ (opow_pos (repr e) omega0_pos) (omega0_le_oadd e n a) /-- Compare ordinal notations -/ def cmp : ONote → ONote → Ordering @@ -231,14 +234,14 @@ theorem NF.zero_of_zero {e n a} (h : NF (ONote.oadd e n a)) (e0 : e = 0) : a = 0 theorem NFBelow.repr_lt {o b} (h : NFBelow o b) : repr o < ω ^ b := by induction h with - | zero => exact opow_pos _ omega_pos + | zero => exact opow_pos _ omega0_pos | oadd' _ _ h₃ _ IH => rw [repr] apply ((add_lt_add_iff_left _).2 IH).trans_le rw [← mul_succ] - apply (mul_le_mul_left' (succ_le_of_lt (nat_lt_omega _)) _).trans + apply (mul_le_mul_left' (succ_le_of_lt (nat_lt_omega0 _)) _).trans rw [← opow_succ] - exact opow_le_opow_right omega_pos (succ_le_of_lt h₃) + exact opow_le_opow_right omega0_pos (succ_le_of_lt h₃) theorem NFBelow.mono {o b₁ b₂} (bb : b₁ ≤ b₂) (h : NFBelow o b₁) : NFBelow o b₂ := by induction h with @@ -253,7 +256,7 @@ theorem NF.below_of_lt' : ∀ {o b}, repr o < ω ^ b → NF o → NFBelow o b | 0, _, _, _ => NFBelow.zero | ONote.oadd _ _ _, _, H, h => h.below_of_lt <| - (opow_lt_opow_iff_right one_lt_omega).1 <| lt_of_le_of_lt (omega_le_oadd _ _ _) H + (opow_lt_opow_iff_right one_lt_omega0).1 <| lt_of_le_of_lt (omega0_le_oadd _ _ _) H theorem nfBelow_ofNat : ∀ n, NFBelow (ofNat n) 1 | 0 => NFBelow.zero @@ -267,13 +270,13 @@ instance nf_one : NF 1 := by rw [← ofNat_one]; infer_instance theorem oadd_lt_oadd_1 {e₁ n₁ o₁ e₂ n₂ o₂} (h₁ : NF (oadd e₁ n₁ o₁)) (h : e₁ < e₂) : oadd e₁ n₁ o₁ < oadd e₂ n₂ o₂ := @lt_of_lt_of_le _ _ (repr (oadd e₁ n₁ o₁)) _ _ - (NF.below_of_lt h h₁).repr_lt (omega_le_oadd e₂ n₂ o₂) + (NF.below_of_lt h h₁).repr_lt (omega0_le_oadd e₂ n₂ o₂) theorem oadd_lt_oadd_2 {e o₁ o₂ : ONote} {n₁ n₂ : ℕ+} (h₁ : NF (oadd e n₁ o₁)) (h : (n₁ : ℕ) < n₂) : 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 _ omega_pos), succ_le_iff, natCast_lt] + rwa [← mul_succ,Ordinal.mul_le_mul_iff_left (opow_pos _ omega0_pos), succ_le_iff, natCast_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 @@ -325,7 +328,7 @@ theorem repr_inj {a b} [NF a] [NF b] : repr a = repr b ↔ a = b := | Ordering.eq, h => h, congr_arg _⟩ -theorem NF.of_dvd_omega_opow {b e n a} (h : NF (ONote.oadd e n a)) +theorem NF.of_dvd_omega0_opow {b e n a} (h : NF (ONote.oadd e n a)) (d : ω ^ b ∣ repr (ONote.oadd e n a)) : b ≤ repr e ∧ ω ^ b ∣ repr a := by have := mt repr_inj.1 (fun h => by injection h : ONote.oadd e n a ≠ 0) @@ -333,9 +336,15 @@ theorem NF.of_dvd_omega_opow {b e n a} (h : NF (ONote.oadd e n a)) simp only [repr] at d exact ⟨L, (dvd_add_iff <| (opow_dvd_opow _ L).mul_right _).1 d⟩ -theorem NF.of_dvd_omega {e n a} (h : NF (ONote.oadd e n a)) : +@[deprecated (since := "2024-09-30")] +alias NF.of_dvd_omega_opow := NF.of_dvd_omega0_opow + +theorem NF.of_dvd_omega0 {e n a} (h : NF (ONote.oadd e n a)) : ω ∣ repr (ONote.oadd e n a) → repr e ≠ 0 ∧ ω ∣ repr a := by - (rw [← opow_one ω, ← one_le_iff_ne_zero]; exact h.of_dvd_omega_opow) + (rw [← opow_one ω, ← one_le_iff_ne_zero]; exact h.of_dvd_omega0_opow) + +@[deprecated (since := "2024-09-30")] +alias NF.of_dvd_omega := NF.of_dvd_omega0 /-- `TopBelow b o` asserts that the largest exponent in `o`, if it exists, is less than `b`. This is an auxiliary definition @@ -443,7 +452,7 @@ theorem repr_add : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ + o₂) = rep unfold repr at this 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') omega_pos).2 + · simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos (repr e') omega0_pos).2 (natCast_le.2 n'.pos) · rw [ee, ← add_assoc, ← mul_add] @@ -503,7 +512,7 @@ theorem repr_sub : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ - o₂) = rep simpa using mul_le_mul_left' (natCast_le.2 <| Nat.succ_pos _) _ · exact (Ordinal.sub_eq_of_add_eq <| - add_absorp (h₂.below_of_lt ee).repr_lt <| omega_le_oadd _ _ _).symm + add_absorp (h₂.below_of_lt ee).repr_lt <| omega0_le_oadd _ _ _).symm /-- Multiplication of ordinal notations (correct only for normal input) -/ def mul : ONote → ONote → ONote @@ -557,7 +566,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 _ omega_pos).2 (natCast_le.2 n₁.2) + simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (natCast_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] @@ -570,8 +579,8 @@ 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 (opow_isLimit_left omega_isLimit this), mul_assoc, - mul_omega_dvd (natCast_pos.2 n₁.pos) (nat_lt_omega _)] + rw [add_mul_limit ao (opow_isLimit_left omega0_isLimit 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) /-- Calculate division and remainder of `o` mod ω. @@ -680,7 +689,7 @@ theorem nf_repr_split' : ∀ {o o' m} [NF o], split' o = (o', m) → NF o' ∧ r · simp at this ⊢ refine IH₁.below_of_lt' - ((Ordinal.mul_lt_mul_iff_left omega_pos).1 <| lt_of_le_of_lt (le_add_right _ m') ?_) + ((Ordinal.mul_lt_mul_iff_left omega0_pos).1 <| lt_of_le_of_lt (le_add_right _ m') ?_) rw [← this, ← IH₂] exact h.snd'.repr_lt · rw [this] @@ -723,9 +732,9 @@ theorem split_dvd {o o' m} [NF o] (h : split o = (o', m)) : ω ∣ repr o' := by theorem split_add_lt {o e n a m} [NF o] (h : split o = (oadd e n a, m)) : repr a + m < ω ^ repr e := by cases' nf_repr_split h with h₁ h₂ - cases' h₁.of_dvd_omega (split_dvd h) with e0 d - apply principal_add_omega_opow _ h₁.snd'.repr_lt (lt_of_lt_of_le (nat_lt_omega _) _) - simpa using opow_le_opow_right omega_pos (one_le_iff_ne_zero.2 e0) + cases' h₁.of_dvd_omega0 (split_dvd h) with e0 d + apply principal_add_omega0_opow _ h₁.snd'.repr_lt (lt_of_lt_of_le (nat_lt_omega0 _) _) + simpa using opow_le_opow_right omega0_pos (one_le_iff_ne_zero.2 e0) @[simp] theorem mulNat_eq_mul (n o) : mulNat o n = o * ofNat n := by cases o <;> cases n <;> rfl @@ -781,22 +790,22 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr (ω ^ repr e) ^ (ω : Ordinal.{0}) := by subst aa have No := Ne.oadd n (Na.below_of_lt' h) - have := omega_le_oadd e n a + 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 _ omega_pos).trans_le this).ne' omega_isLimit).2 + apply (opow_le_of_limit ((opow_pos _ omega0_pos).trans_le this).ne' omega0_isLimit).2 intro b l have := (No.below_of_lt (lt_succ _)).repr_lt rw [repr] at this apply (opow_le_opow_left b <| this.le).trans rw [← opow_mul, ← opow_mul] - apply opow_le_opow_right omega_pos + apply opow_le_opow_right omega0_pos rcases le_or_lt ω (repr e) with h | h · apply (mul_le_mul_left' (le_succ b) _).trans - rw [← add_one_eq_succ, add_mul_succ _ (one_add_of_omega_le h), add_one_eq_succ, succ_le_iff, + 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 omega_isLimit.2 _ l - · apply (principal_mul_omega (omega_isLimit.2 _ h) l).le.trans + exact omega0_isLimit.2 _ l + · apply (principal_mul_omega0 (omega0_isLimit.2 _ h) l).le.trans simpa using mul_le_mul_right' (one_le_iff_ne_zero.2 e0) ω section @@ -827,30 +836,30 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω · simp only [R', ONote.repr_scale, ONote.repr, ONote.mulNat_eq_mul, ONote.opowAux, ONote.repr_ofNat, ONote.repr_mul, ONote.repr_add, Ordinal.opow_mul, ONote.zero_add] have α0 : 0 < α' := by simpa [lt_def, repr] using oadd_pos a0 n a' - have ω00 : 0 < ω0 ^ (k : Ordinal) := opow_pos _ (opow_pos _ omega_pos) + have ω00 : 0 < ω0 ^ (k : Ordinal) := opow_pos _ (opow_pos _ omega0_pos) have Rl : R < ω ^ (repr a0 * succ ↑k) := by by_cases k0 : k = 0 · simp only [k0, Nat.cast_zero, succ_zero, mul_one, R] - refine lt_of_lt_of_le ?_ (opow_le_opow_right omega_pos (one_le_iff_ne_zero.2 e0)) - cases' m with m <;> simp [opowAux, omega_pos] + refine lt_of_lt_of_le ?_ (opow_le_opow_right omega0_pos (one_le_iff_ne_zero.2 e0)) + cases' m with m <;> simp [opowAux, omega0_pos] rw [← add_one_eq_succ, ← Nat.cast_succ] - apply nat_lt_omega + apply nat_lt_omega0 · rw [opow_mul] exact IH.1 k0 refine ⟨fun _ => ?_, ?_⟩ · rw [RR, ← opow_mul _ _ (succ k.succ)] have e0 := Ordinal.pos_iff_ne_zero.2 e0 have rr0 : 0 < repr a0 + repr a0 := lt_of_lt_of_le e0 (le_add_left _ _) - apply principal_add_omega_opow + apply principal_add_omega0_opow · simp only [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one, add_one_eq_succ, opow_mul, opow_succ, mul_assoc] rw [Ordinal.mul_lt_mul_iff_left ω00, ← Ordinal.opow_add] have : _ < ω ^ (repr a0 + repr a0) := (No.below_of_lt ?_).repr_lt - · exact mul_lt_omega_opow rr0 this (nat_lt_omega _) + · exact mul_lt_omega0_opow rr0 this (nat_lt_omega0 _) · simpa using (add_lt_add_iff_left (repr a0)).2 e0 · exact lt_of_lt_of_le Rl - (opow_le_opow_right omega_pos <| + (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))) _) calc (ω0 ^ (k.succ : Ordinal)) * α' + R' @@ -862,10 +871,10 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω · have α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_omega_dvd.2 ⟨ne_of_gt α0, αd⟩), mul_assoc, - @mul_omega_dvd n (natCast_pos.2 n.pos) (nat_lt_omega _) _ αd] + 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] apply @add_absorp _ (repr a0 * succ ↑k) - · refine principal_add_omega_opow _ ?_ Rl + · 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)) @@ -878,7 +887,7 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω apply add_absorp Rl rw [opow_mul, opow_succ] apply mul_le_mul_left' - simpa [repr] using omega_le_oadd a0 n a' + simpa [repr] using omega0_le_oadd a0 n a' end @@ -897,18 +906,18 @@ theorem repr_opow (o₁ o₂) [NF o₁] [NF o₂] : repr (o₁ ^ o₂) = repr o simp only [opow_def, opowAux2, opow, e₁, h, r₁, e₂, r₂, repr, opow_zero, Nat.succPNat_coe, Nat.cast_succ, Nat.cast_zero, _root_.zero_add, mul_one, add_zero, one_opow, npow_eq_pow] - rw [opow_add, opow_mul, opow_omega, add_one_eq_succ] + rw [opow_add, opow_mul, opow_omega0, add_one_eq_succ] · congr conv_lhs => dsimp [(· ^ ·)] simp [Pow.pow, opow, Ordinal.succ_ne_zero] rw [opow_natCast] · simpa [Nat.one_le_iff_ne_zero] - · rw [← Nat.cast_succ, lt_omega] + · rw [← Nat.cast_succ, lt_omega0] exact ⟨_, rfl⟩ · haveI := N₁.fst haveI := N₁.snd - cases' N₁.of_dvd_omega (split_dvd e₁) with a00 ad + cases' N₁.of_dvd_omega0 (split_dvd e₁) with a00 ad have al := split_add_lt e₁ have aa : repr (a' + ofNat m) = repr a' + m := by simp only [eq_self_iff_true, ONote.repr_ofNat, ONote.repr_add] @@ -955,13 +964,13 @@ private theorem exists_lt_add {α} [hα : Nonempty α] {o : Ordinal} {f : α → refine (H h).imp fun i H => ?_ rwa [← Ordinal.add_sub_cancel_of_le h', add_lt_add_iff_left] -private theorem exists_lt_mul_omega' {o : Ordinal} ⦃a⦄ (h : a < o * ω) : +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 omega_isLimit).1 h - obtain ⟨i, rfl⟩ := lt_omega.1 hi + obtain ⟨i, hi, h'⟩ := (lt_mul_of_limit omega0_isLimit).1 h + obtain ⟨i, rfl⟩ := lt_omega0.1 hi exact ⟨i, h'.trans_le (le_add_right _ _)⟩ -private theorem exists_lt_omega_opow' {α} {o b : Ordinal} (hb : 1 < b) (ho : o.IsLimit) +private theorem exists_lt_omega0_opow' {α} {o b : Ordinal} (hb : 1 < b) (ho : o.IsLimit) {f : α → Ordinal} (H : ∀ ⦃a⦄, a < o → ∃ i, a < f i) ⦃a⦄ (h : a < b ^ o) : ∃ i, a < b ^ f i := by obtain ⟨d, hd, h'⟩ := (lt_opow_of_limit (zero_lt_one.trans hb).ne' ho).1 h @@ -1014,39 +1023,40 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta have := PNat.natPred_add_one m; rw [e'] at this; exact PNat.coe_inj.1 this.symm]) <;> (try rw [show m = (m' + 1).succPNat by rw [← e', ← PNat.coe_inj, Nat.succPNat_coe, ← Nat.add_one, PNat.natPred_add_one]]) <;> - simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, - eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, - Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, + simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega0, add_lt_add_iff_left, + add_zero, eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, + Nat.cast_succ, Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, _root_.zero_add, zero_def] · decide · exact ⟨rfl, inferInstance⟩ - · have := opow_pos (repr a') omega_pos + · have := opow_pos (repr a') omega0_pos refine - ⟨mul_isLimit this omega_isLimit, fun i => - ⟨this, ?_, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega'⟩ + ⟨mul_isLimit this omega0_isLimit, 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_omega - · have := opow_pos (repr a') omega_pos + apply nat_lt_omega0 + · have := opow_pos (repr a') omega0_pos refine - ⟨add_isLimit _ (mul_isLimit this omega_isLimit), fun i => ⟨this, ?_, ?_⟩, - exists_lt_add exists_lt_mul_omega'⟩ + ⟨add_isLimit _ (mul_isLimit this omega0_isLimit), 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_omega + apply nat_lt_omega0 · refine fun H => H.fst.oadd _ (NF.below_of_lt' ?_ (@NF.oadd_zero _ _ (iha.2 H.fst))) rw [repr, ← zero_def, repr, add_zero, iha.1, opow_succ, Ordinal.mul_lt_mul_iff_left this] - apply nat_lt_omega + apply nat_lt_omega0 · rcases iha with ⟨h1, h2, h3⟩ - refine ⟨opow_isLimit one_lt_omega h1, fun i => ?_, exists_lt_omega_opow' one_lt_omega h1 h3⟩ + refine ⟨opow_isLimit one_lt_omega0 h1, fun i => ?_, + exists_lt_omega0_opow' one_lt_omega0 h1 h3⟩ obtain ⟨h4, h5, h6⟩ := h2 i exact ⟨h4, h5, fun H => @NF.oadd_zero _ _ (h6 H.fst)⟩ · rcases iha with ⟨h1, h2, h3⟩ refine - ⟨add_isLimit _ (opow_isLimit one_lt_omega h1), fun i => ?_, - exists_lt_add (exists_lt_omega_opow' one_lt_omega h1 h3)⟩ + ⟨add_isLimit _ (opow_isLimit one_lt_omega0 h1), fun i => ?_, + exists_lt_add (exists_lt_omega0_opow' one_lt_omega0 h1 h3)⟩ obtain ⟨h4, h5, h6⟩ := h2 i refine ⟨h4, h5, fun H => H.fst.oadd _ (NF.below_of_lt' ?_ (@NF.oadd_zero _ _ (h6 H.fst)))⟩ rwa [repr, ← zero_def, repr, add_zero, PNat.one_coe, Nat.cast_one, mul_one, - opow_lt_opow_iff_right one_lt_omega] + opow_lt_opow_iff_right one_lt_omega0] · refine ⟨by rw [repr, ihb.1, add_succ, repr], fun H => H.fst.oadd _ (NF.below_of_lt' ?_ (ihb.2 H.snd))⟩ have := H.snd'.repr_lt diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index f149eea4b37a4..2835fd286b3f3 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -14,9 +14,9 @@ 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. -* `principal_add_iff_zero_or_omega_opow`: The main characterization theorem for additive principal +* `principal_add_iff_zero_or_omega0_opow`: The main characterization theorem for additive principal ordinals. -* `principal_mul_iff_le_two_or_omega_opow_opow`: The main characterization theorem for +* `principal_mul_iff_le_two_or_omega0_opow_opow`: The main characterization theorem for multiplicative principal ordinals. ## TODO @@ -163,39 +163,51 @@ theorem principal_add_iff_add_lt_ne_self {a} : rcases exists_lt_add_of_not_principal_add ha with ⟨b, hb, c, hc, rfl⟩ exact (H b hb c hc).irrefl⟩ -theorem add_omega {a : Ordinal} (h : a < ω) : a + ω = ω := by - rcases lt_omega.1 h with ⟨n, rfl⟩ +theorem add_omega0 {a : Ordinal} (h : a < ω) : a + ω = ω := by + rcases lt_omega0.1 h with ⟨n, rfl⟩ clear h; induction' n with n IH · rw [Nat.cast_zero, zero_add] - · rwa [Nat.cast_succ, add_assoc, one_add_of_omega_le (le_refl _)] + · rwa [Nat.cast_succ, add_assoc, one_add_of_omega0_le (le_refl _)] -theorem principal_add_omega : Principal (· + ·) ω := - principal_add_iff_add_left_eq_self.2 fun _ => add_omega +@[deprecated (since := "2024-09-30")] +alias add_omega := add_omega0 -theorem add_omega_opow {a b : Ordinal} (h : a < ω ^ b) : a + ω ^ b = ω ^ b := by +theorem principal_add_omega0 : Principal (· + ·) ω := + principal_add_iff_add_left_eq_self.2 fun _ => add_omega0 + +@[deprecated (since := "2024-09-30")] +alias principal_add_omega := principal_add_omega0 + +theorem add_omega0_opow {a b : Ordinal} (h : a < ω ^ b) : a + ω ^ b = ω ^ b := by refine le_antisymm ?_ (le_add_left _ a) induction' b using limitRecOn with b _ b l IH · 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 omega_isLimit).1 h with ⟨x, xo, ax⟩ + rcases (lt_mul_of_limit omega0_isLimit).1 h with ⟨x, xo, ax⟩ apply (add_le_add_right ax.le _).trans - rw [opow_succ, ← mul_add, add_omega xo] - · rcases (lt_opow_of_limit omega_ne_zero l).1 h with ⟨x, xb, ax⟩ - apply (((add_isNormal a).trans <| opow_isNormal one_lt_omega).limit_le l).2 + rw [opow_succ, ← mul_add, add_omega0 xo] + · rcases (lt_opow_of_limit omega0_ne_zero l).1 h with ⟨x, xb, ax⟩ + apply (((add_isNormal a).trans <| opow_isNormal one_lt_omega0).limit_le l).2 intro y yb calc a + ω ^ y ≤ a + ω ^ max x y := - add_le_add_left (opow_le_opow_right omega_pos (le_max_right x y)) _ + add_le_add_left (opow_le_opow_right omega0_pos (le_max_right x y)) _ _ ≤ ω ^ max x y := - IH _ (max_lt xb yb) <| ax.trans_le <| opow_le_opow_right omega_pos <| le_max_left x y + IH _ (max_lt xb yb) <| ax.trans_le <| opow_le_opow_right omega0_pos <| le_max_left x y _ ≤ ω ^ b := - opow_le_opow_right omega_pos <| (max_lt xb yb).le + opow_le_opow_right omega0_pos <| (max_lt xb yb).le + +@[deprecated (since := "2024-09-30")] +alias add_omega_opow := add_omega0_opow + +theorem principal_add_omega0_opow (o : Ordinal) : Principal (· + ·) (ω ^ o) := + principal_add_iff_add_left_eq_self.2 fun _ => add_omega0_opow -theorem principal_add_omega_opow (o : Ordinal) : Principal (· + ·) (ω ^ o) := - principal_add_iff_add_left_eq_self.2 fun _ => add_omega_opow +@[deprecated (since := "2024-09-30")] +alias principal_add_omega_opow := principal_add_omega0_opow /-- The main characterization theorem for additive principal ordinals. -/ -theorem principal_add_iff_zero_or_omega_opow {o : Ordinal} : +theorem principal_add_iff_zero_or_omega0_opow {o : Ordinal} : Principal (· + ·) o ↔ o = 0 ∨ o ∈ Set.range (ω ^ · : Ordinal → Ordinal) := by rcases eq_or_ne o 0 with (rfl | ho) · simp only [principal_zero, Or.inl] @@ -203,12 +215,12 @@ theorem principal_add_iff_zero_or_omega_opow {o : Ordinal} : simp only [ho, false_or] refine ⟨fun H => ⟨_, ((lt_or_eq_of_le (opow_log_le_self _ ho)).resolve_left fun h => ?_)⟩, - fun ⟨b, e⟩ => e.symm ▸ fun a => add_omega_opow⟩ + fun ⟨b, e⟩ => e.symm ▸ fun a => add_omega0_opow⟩ have := H _ h - have := lt_opow_succ_log_self one_lt_omega o - rw [opow_succ, lt_mul_of_limit omega_isLimit] at this + have := lt_opow_succ_log_self one_lt_omega0 o + rw [opow_succ, lt_mul_of_limit omega0_isLimit] at this rcases this with ⟨a, ao, h'⟩ - rcases lt_omega.1 ao with ⟨n, rfl⟩ + rcases lt_omega0.1 ao with ⟨n, rfl⟩ clear ao revert h' apply not_lt_of_le @@ -218,18 +230,21 @@ theorem principal_add_iff_zero_or_omega_opow {o : Ordinal} : · simp [Nat.cast_zero, mul_zero, zero_add] · simp only [Nat.cast_succ, mul_add_one, add_assoc, this, IH] +@[deprecated (since := "2024-09-30")] +alias principal_add_iff_zero_or_omega_opow := principal_add_iff_zero_or_omega0_opow + theorem opow_principal_add_of_principal_add {a} (ha : Principal (· + ·) a) (b : Ordinal) : Principal (· + ·) (a ^ b) := by - rcases principal_add_iff_zero_or_omega_opow.1 ha with (rfl | ⟨c, rfl⟩) + rcases principal_add_iff_zero_or_omega0_opow.1 ha with (rfl | ⟨c, rfl⟩) · rcases eq_or_ne b 0 with (rfl | hb) · rw [opow_zero] exact principal_add_one · rwa [zero_opow hb] · rw [← opow_mul] - exact principal_add_omega_opow _ + exact principal_add_omega0_opow _ theorem add_absorp {a b c : Ordinal} (h₁ : a < ω ^ b) (h₂ : ω ^ b ≤ c) : a + c = c := by - rw [← Ordinal.add_sub_cancel_of_le h₂, ← add_assoc, add_omega_opow h₁] + rw [← Ordinal.add_sub_cancel_of_le h₂, ← add_assoc, add_omega0_opow h₁] theorem mul_principal_add_is_principal_add (a : Ordinal.{u}) {b : Ordinal.{u}} (hb₁ : b ≠ 1) (hb : Principal (· + ·) b) : Principal (· + ·) (a * b) := by @@ -302,44 +317,59 @@ theorem principal_mul_iff_mul_left_eq {o : Ordinal} : rw [← h a ha hao] exact (mul_isNormal ha).strictMono hbo -theorem principal_mul_omega : Principal (· * ·) ω := fun a b ha hb => - match a, b, lt_omega.1 ha, lt_omega.1 hb with +theorem principal_mul_omega0 : Principal (· * ·) ω := fun a b ha hb => + match a, b, lt_omega0.1 ha, lt_omega0.1 hb with | _, _, ⟨m, rfl⟩, ⟨n, rfl⟩ => by dsimp only; rw [← natCast_mul] - apply nat_lt_omega + apply nat_lt_omega0 + +@[deprecated (since := "2024-09-30")] +alias principal_mul_omega := principal_mul_omega0 + +theorem mul_omega0 {a : Ordinal} (a0 : 0 < a) (ha : a < ω) : a * ω = ω := + principal_mul_iff_mul_left_eq.1 principal_mul_omega0 a a0 ha -theorem mul_omega {a : Ordinal} (a0 : 0 < a) (ha : a < ω) : a * ω = ω := - principal_mul_iff_mul_left_eq.1 principal_mul_omega a a0 ha +@[deprecated (since := "2024-09-30")] +alias mul_omega := mul_omega0 -theorem mul_lt_omega_opow {a b c : Ordinal} (c0 : 0 < c) (ha : a < ω ^ c) (hb : b < ω) : +theorem mul_lt_omega0_opow {a b c : Ordinal} (c0 : 0 < c) (ha : a < ω ^ c) (hb : b < ω) : a * b < ω ^ c := by rcases zero_or_succ_or_limit c with (rfl | ⟨c, rfl⟩ | l) · exact (lt_irrefl _).elim c0 · rw [opow_succ] at ha - rcases ((mul_isNormal <| opow_pos _ omega_pos).limit_lt omega_isLimit).1 ha with ⟨n, hn, an⟩ + rcases ((mul_isNormal <| opow_pos _ omega0_pos).limit_lt omega0_isLimit).1 ha with ⟨n, hn, an⟩ apply (mul_le_mul_right' (le_of_lt an) _).trans_lt - rw [opow_succ, mul_assoc, mul_lt_mul_iff_left (opow_pos _ omega_pos)] - exact principal_mul_omega hn hb - · rcases ((opow_isNormal one_lt_omega).limit_lt l).1 ha with ⟨x, hx, ax⟩ + rw [opow_succ, mul_assoc, mul_lt_mul_iff_left (opow_pos _ omega0_pos)] + exact principal_mul_omega0 hn hb + · rcases ((opow_isNormal one_lt_omega0).limit_lt l).1 ha with ⟨x, hx, ax⟩ refine (mul_le_mul' (le_of_lt ax) (le_of_lt hb)).trans_lt ?_ - rw [← opow_succ, opow_lt_opow_iff_right one_lt_omega] + rw [← opow_succ, opow_lt_opow_iff_right one_lt_omega0] exact l.2 _ hx -theorem mul_omega_opow_opow {a b : Ordinal} (a0 : 0 < a) (h : a < ω ^ ω ^ b) : +@[deprecated (since := "2024-09-30")] +alias mul_lt_omega_opow := mul_lt_omega0_opow + +theorem mul_omega0_opow_opow {a b : Ordinal} (a0 : 0 < a) (h : a < ω ^ ω ^ b) : a * ω ^ ω ^ b = ω ^ ω ^ b := by obtain rfl | b0 := eq_or_ne b 0 · rw [opow_zero, opow_one] at h ⊢ - exact mul_omega a0 h + exact mul_omega0 a0 h · apply le_antisymm · obtain ⟨x, xb, ax⟩ := - (lt_opow_of_limit omega_ne_zero (opow_isLimit_left omega_isLimit b0)).1 h + (lt_opow_of_limit omega0_ne_zero (opow_isLimit_left omega0_isLimit b0)).1 h apply (mul_le_mul_right' (le_of_lt ax) _).trans - rw [← opow_add, add_omega_opow xb] + rw [← opow_add, add_omega0_opow xb] · conv_lhs => rw [← one_mul (ω ^ _)] exact mul_le_mul_right' (one_le_iff_pos.2 a0) _ -theorem principal_mul_omega_opow_opow (o : Ordinal) : Principal (· * ·) (ω ^ ω ^ o) := - principal_mul_iff_mul_left_eq.2 fun _ => mul_omega_opow_opow +@[deprecated (since := "2024-09-30")] +alias mul_omega_opow_opow := mul_omega0_opow_opow + +theorem principal_mul_omega0_opow_opow (o : Ordinal) : Principal (· * ·) (ω ^ ω ^ o) := + principal_mul_iff_mul_left_eq.2 fun _ => mul_omega0_opow_opow + +@[deprecated (since := "2024-09-30")] +alias principal_mul_omega_opow_opow := principal_mul_omega0_opow_opow theorem principal_add_of_principal_mul_opow {o b : Ordinal} (hb : 1 < b) (ho : Principal (· * ·) (b ^ o)) : Principal (· + ·) o := by @@ -349,24 +379,30 @@ theorem principal_add_of_principal_mul_opow {o b : Ordinal} (hb : 1 < b) rwa [← opow_add, opow_lt_opow_iff_right hb] at this /-- The main characterization theorem for multiplicative principal ordinals. -/ -theorem principal_mul_iff_le_two_or_omega_opow_opow {o : Ordinal} : +theorem principal_mul_iff_le_two_or_omega0_opow_opow {o : Ordinal} : Principal (· * ·) o ↔ o ≤ 2 ∨ o ∈ Set.range (ω ^ ω ^ · : Ordinal → Ordinal) := by refine ⟨fun ho => ?_, ?_⟩ · rcases le_or_lt o 2 with ho₂ | ho₂ · exact Or.inl ho₂ - · rcases principal_add_iff_zero_or_omega_opow.1 (principal_add_of_principal_mul ho ho₂.ne') with - (rfl | ⟨a, rfl⟩) + · rcases principal_add_iff_zero_or_omega0_opow.1 (principal_add_of_principal_mul ho ho₂.ne') + with (rfl | ⟨a, rfl⟩) · exact (Ordinal.not_lt_zero 2 ho₂).elim - · rcases principal_add_iff_zero_or_omega_opow.1 - (principal_add_of_principal_mul_opow one_lt_omega ho) with (rfl | ⟨b, rfl⟩) + · rcases principal_add_iff_zero_or_omega0_opow.1 + (principal_add_of_principal_mul_opow one_lt_omega0 ho) with (rfl | ⟨b, rfl⟩) · simp · exact Or.inr ⟨b, rfl⟩ · rintro (ho₂ | ⟨a, rfl⟩) · exact principal_mul_of_le_two ho₂ - · exact principal_mul_omega_opow_opow a + · exact principal_mul_omega0_opow_opow a + +@[deprecated (since := "2024-09-30")] +alias principal_mul_iff_le_two_or_omega_opow_opow := principal_mul_iff_le_two_or_omega0_opow_opow -theorem mul_omega_dvd {a : Ordinal} (a0 : 0 < a) (ha : a < ω) : ∀ {b}, ω ∣ b → a * b = b - | _, ⟨b, rfl⟩ => by rw [← mul_assoc, mul_omega a0 ha] +theorem mul_omega0_dvd {a : Ordinal} (a0 : 0 < a) (ha : a < ω) : ∀ {b}, ω ∣ b → a * b = b + | _, ⟨b, rfl⟩ => by rw [← mul_assoc, mul_omega0 a0 ha] + +@[deprecated (since := "2024-09-30")] +alias mul_omega_dvd := mul_omega0_dvd theorem mul_eq_opow_log_succ {a b : Ordinal.{u}} (ha : a ≠ 0) (hb : Principal (· * ·) b) (hb₂ : 2 < b) : a * b = b ^ succ (log b a) := by @@ -387,15 +423,21 @@ theorem mul_eq_opow_log_succ {a b : Ordinal.{u}} (ha : a ≠ 0) (hb : Principal /-! #### Exponential principal ordinals -/ -theorem principal_opow_omega : Principal (· ^ ·) ω := fun a b ha hb => - match a, b, lt_omega.1 ha, lt_omega.1 hb with +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 simp_rw [← natCast_opow] - apply nat_lt_omega + apply nat_lt_omega0 + +@[deprecated (since := "2024-09-30")] +alias principal_opow_omega := principal_opow_omega0 -theorem opow_omega {a : Ordinal} (a1 : 1 < a) (h : a < ω) : a ^ ω = ω := - ((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) omega_isLimit).2 fun _ hb => - (principal_opow_omega h hb).le).antisymm +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 => + (principal_opow_omega0 h hb).le).antisymm (right_le_opow _ a1) +@[deprecated (since := "2024-09-30")] +alias opow_omega := opow_omega0 + end Ordinal diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index d66cd78fe7c09..61a8f1beba401 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3514,8 +3514,8 @@ one_ne_zero' OnePoint.continuousAt_infty' OnePoint.isOpen_iff_of_mem' OnePoint.tendsto_nhds_infty' -ONote.exists_lt_mul_omega' -ONote.exists_lt_omega_opow' +ONote.exists_lt_mul_omega0' +ONote.exists_lt_omega0_opow' ONote.fastGrowing_zero' ONote.NF.below_of_lt' ONote.nf_repr_split' From e95b1d9ccaafa1eeb604d52f0e121d457615bc2b Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 3 Oct 2024 10:05:47 +0000 Subject: [PATCH 146/174] chore: update Mathlib dependencies 2024-10-03 (#17370) 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 1c5880467526f..efc4091d6b8c8 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4756e0fc48acce0cc808df0ad149de5973240df6", + "rev": "34e690ec07f6f6375668adba5a16d0d723226c2c", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 3241ff343767d68e9be2a7b7ed530aa4049f206d Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 3 Oct 2024 10:05:48 +0000 Subject: [PATCH 147/174] chore(Algebra/DirectLimit): rm two porting notes (#17371) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/DirectLimit.lean | 12 +++--------- 1 file changed, 3 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index f1aa6c7623464..f077a686dc355 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -316,12 +316,8 @@ theorem of.zero_exact_aux [∀ i (k : G i), Decidable (k ≠ 0)] [Nonempty ι] [ ⟨k, fun l hl => (Finset.mem_union.1 (DFinsupp.support_add hl)).elim (fun hl => le_trans (hi _ hl) hik) fun hl => le_trans (hj _ hl) hjk, by - -- Porting note: this had been - -- simp [LinearMap.map_add, hxi, hyj, toModule_totalize_of_le hik hi, - -- toModule_totalize_of_le hjk hj] - simp only [map_add] - rw [toModule_totalize_of_le hik hi, toModule_totalize_of_le hjk hj] - simp [hxi, hyj]⟩) + simp [LinearMap.map_add, hxi, hyj, toModule_totalize_of_le hik hi, + toModule_totalize_of_le hjk hj]⟩) fun a x ⟨i, hi, hxi⟩ => ⟨i, fun k hk => hi k (DirectSum.support_smul _ _ hk), by simp [LinearMap.map_smul, hxi]⟩ @@ -721,9 +717,7 @@ theorem of.zero_exact_aux [Nonempty ι] [IsDirected ι (· ≤ ·)] {x : FreeCom dsimp only rw [(f' i i _).map_mul] · exact sub_self _ - all_goals tauto - -- Porting note: was - --exacts [sub_self _, Or.inl rfl, Or.inr (Or.inr rfl), Or.inr (Or.inl rfl)] + exacts [Or.inl rfl, Or.inr (Or.inr rfl), Or.inr (Or.inl rfl)] · refine Nonempty.elim (by infer_instance) fun ind : ι => ?_ refine ⟨ind, ∅, fun _ => False.elim, isSupported_zero, fun [_] => ?_⟩ -- Porting note: `RingHom.map_zero` was `(restriction _).map_zero` From 70102953aa125ae9abd08cd7d0ab8f4194055834 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Thu, 3 Oct 2024 10:05:49 +0000 Subject: [PATCH 148/174] refactor(Counterexamples/SeminormLatticeNotDistrib): minor golfing (#17372) --- Counterexamples/SeminormLatticeNotDistrib.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Counterexamples/SeminormLatticeNotDistrib.lean b/Counterexamples/SeminormLatticeNotDistrib.lean index 770184c227ba7..e547175d5a7d5 100644 --- a/Counterexamples/SeminormLatticeNotDistrib.lean +++ b/Counterexamples/SeminormLatticeNotDistrib.lean @@ -57,14 +57,14 @@ theorem not_distrib : ¬(p ⊔ q1) ⊓ (p ⊔ q2) ≤ p ⊔ q1 ⊓ q2 := by 4 / 3 = 4 * (1 - 2 / 3) := by norm_num _ ≤ 4 * (1 - x.snd) := by gcongr _ ≤ 4 * |1 - x.snd| := by gcongr; apply le_abs_self - _ = q2 ((1, 1) - x) := by simp; rfl + _ = q2 ((1, 1) - x) := rfl _ ≤ (p ⊔ q2) ((1, 1) - x) := le_sup_right _ ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) := le_add_of_nonneg_left (apply_nonneg _ _) · calc 4 / 3 = 2 / 3 + (1 - 1 / 3) := by norm_num _ ≤ x.snd + (1 - x.fst) := by gcongr _ ≤ |x.snd| + |1 - x.fst| := add_le_add (le_abs_self _) (le_abs_self _) - _ ≤ p x + p ((1, 1) - x) := by exact add_le_add le_sup_right le_sup_left + _ ≤ p x + p ((1, 1) - x) := add_le_add le_sup_right le_sup_left _ ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) := add_le_add le_sup_left le_sup_left · calc 4 / 3 = 4 * (1 / 3) := by norm_num From 81d4159870a16f4ae3253172f5d203bfa004b8e0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 10:58:26 +0000 Subject: [PATCH 149/174] chore: generalise more lemmas from `LinearOrderedField` to `GroupWithZero` (#17359) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and move the lemmas from `Algebra.Order.Field.Basic` to `Algebra.Order.GroupWithZero.Unbundled`. Also take the opportunity to add `₀` at the end, per the naming convention, and to make the priming of the lemmas consistent among pairs (primes should be on the lemmas changing left multiplication into right multiplication). --- Archive/Imo/Imo1972Q5.lean | 4 +- Archive/Imo/Imo1986Q5.lean | 2 +- Mathlib.lean | 1 + .../Computation/ApproximationCorollaries.lean | 4 +- Mathlib/Algebra/Order/Archimedean/Basic.lean | 10 +- Mathlib/Algebra/Order/Field/Basic.lean | 82 +++++------ Mathlib/Algebra/Order/Field/Pointwise.lean | 12 +- Mathlib/Algebra/Order/Floor.lean | 12 +- .../Order/GroupWithZero/Canonical.lean | 6 - .../Order/GroupWithZero/Unbundled.lean | 137 +++++++++++++++++- .../Order/GroupWithZero/Unbundled/Lemmas.lean | 27 ++++ Mathlib/Analysis/Asymptotics/Asymptotics.lean | 4 +- .../Analysis/CStarAlgebra/Unitization.lean | 2 +- .../ApproximatesLinearOn.lean | 2 +- Mathlib/Analysis/Calculus/MeanValue.lean | 2 +- .../Analysis/Calculus/UniformLimitsDeriv.lean | 4 +- Mathlib/Analysis/Complex/AbelLimit.lean | 2 +- Mathlib/Analysis/Complex/Hadamard.lean | 6 +- .../Analysis/Complex/PhragmenLindelof.lean | 2 +- .../Complex/UpperHalfPlane/Metric.lean | 2 +- Mathlib/Analysis/Convex/Body.lean | 2 +- Mathlib/Analysis/Convex/Continuous.lean | 2 +- Mathlib/Analysis/Convex/Deriv.lean | 4 +- Mathlib/Analysis/Convex/Gauge.lean | 8 +- Mathlib/Analysis/Convex/Integral.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 4 +- .../Convex/SpecificFunctions/Basic.lean | 33 +++-- .../Analysis/Convex/StrictConvexSpace.lean | 2 +- .../Analysis/Distribution/SchwartzSpace.lean | 4 +- .../Fourier/RiemannLebesgueLemma.lean | 8 +- .../InnerProductSpace/LaxMilgram.lean | 2 +- .../Analysis/LocallyConvex/WithSeminorms.lean | 2 +- Mathlib/Analysis/Normed/Affine/AddTorsor.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- .../Analysis/Normed/Field/ProperSpace.lean | 2 +- Mathlib/Analysis/Normed/Group/Quotient.lean | 4 +- Mathlib/Analysis/Normed/Module/Basic.lean | 2 +- .../NormedSpace/OperatorNorm/NNNorm.lean | 2 +- Mathlib/Analysis/NormedSpace/Pointwise.lean | 2 +- Mathlib/Analysis/NormedSpace/RieszLemma.lean | 6 +- Mathlib/Analysis/Seminorm.lean | 12 +- .../SpecialFunctions/Complex/Arctan.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 2 +- .../SpecialFunctions/JapaneseBracket.lean | 2 +- .../Analysis/SpecialFunctions/Log/Base.lean | 2 +- .../Analysis/SpecialFunctions/Log/Basic.lean | 2 +- .../SpecialFunctions/Pow/Complex.lean | 2 +- .../Analysis/SpecialFunctions/Pow/Real.lean | 2 +- .../SpecialFunctions/Trigonometric/Angle.lean | 4 +- .../Trigonometric/Arctan.lean | 2 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 2 +- .../Additive/AP/Three/Behrend.lean | 4 +- Mathlib/Combinatorics/Schnirelmann.lean | 2 +- .../SimpleGraph/Regularity/Bound.lean | 4 +- .../SimpleGraph/Regularity/Uniform.lean | 2 +- Mathlib/Data/Complex/Exponential.lean | 4 +- Mathlib/Data/Int/CardIntervalMod.lean | 10 +- Mathlib/Data/NNReal/Basic.lean | 29 ++-- Mathlib/Data/Real/Archimedean.lean | 2 +- Mathlib/Data/Real/Pi/Bounds.lean | 4 +- Mathlib/Data/Real/Pi/Irrational.lean | 4 +- Mathlib/Data/Set/Pointwise/Interval.lean | 8 +- .../Geometry/Manifold/Instances/Sphere.lean | 2 +- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Gershgorin.lean | 2 +- .../Covering/DensityTheorem.lean | 2 +- Mathlib/MeasureTheory/Covering/Vitali.lean | 2 +- .../Function/ConvergenceInMeasure.lean | 2 +- .../Integral/MeanInequalities.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 2 +- .../Measure/Lebesgue/VolumeOfBalls.lean | 4 +- .../ClassNumber/AdmissibleAbs.lean | 4 +- .../ClassNumber/AdmissibleCardPowDegree.lean | 2 +- .../DiophantineApproximation.lean | 8 +- Mathlib/NumberTheory/Liouville/Basic.lean | 4 +- .../NumberTheory/Liouville/LiouvilleWith.lean | 4 +- Mathlib/NumberTheory/Modular.lean | 4 +- .../CanonicalEmbedding/ConvexBody.lean | 2 +- .../NumberTheory/NumberField/ClassNumber.lean | 2 +- .../NumberField/Discriminant.lean | 4 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- Mathlib/NumberTheory/Rayleigh.lean | 8 +- Mathlib/RingTheory/RootsOfUnity/Complex.lean | 2 +- .../Algebra/Module/FiniteDimension.lean | 2 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 4 +- Mathlib/Topology/ContinuousMap/Ideals.lean | 2 +- 87 files changed, 373 insertions(+), 228 deletions(-) create mode 100644 Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean diff --git a/Archive/Imo/Imo1972Q5.lean b/Archive/Imo/Imo1972Q5.lean index f96965f9d6a75..ba2e4d8c6a0a0 100644 --- a/Archive/Imo/Imo1972Q5.lean +++ b/Archive/Imo/Imo1972Q5.lean @@ -50,7 +50,7 @@ theorem imo1972_q5 (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - y calc 0 < ‖f x‖ := norm_pos_iff.mpr hx _ ≤ k := hk₁ x - rw [div_lt_iff] + rw [div_lt_iff₀] · apply lt_mul_of_one_lt_right h₁ hneg · exact zero_lt_one.trans hneg -- Demonstrate that `k ≤ k'` using `hk₂`. @@ -87,7 +87,7 @@ theorem imo1972_q5' (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - have h : ∀ x, ‖f x‖ ≤ k := le_ciSup hf2 have hgy : 0 < ‖g y‖ := by linarith have k_pos : 0 < k := lt_of_lt_of_le (norm_pos_iff.mpr hx) (h x) - have : k / ‖g y‖ < k := (div_lt_iff hgy).mpr (lt_mul_of_one_lt_right k_pos H) + have : k / ‖g y‖ < k := (div_lt_iff₀ hgy).mpr (lt_mul_of_one_lt_right k_pos H) have : k ≤ k / ‖g y‖ := by suffices ∀ x, ‖f x‖ ≤ k / ‖g y‖ from ciSup_le this intro x diff --git a/Archive/Imo/Imo1986Q5.lean b/Archive/Imo/Imo1986Q5.lean index 6789efcbb837d..82fe5961c1647 100644 --- a/Archive/Imo/Imo1986Q5.lean +++ b/Archive/Imo/Imo1986Q5.lean @@ -54,7 +54,7 @@ theorem map_of_lt_two (hx : x < 2) : f x = 2 / (2 - x) := by have hx' : 0 < 2 - x := tsub_pos_of_lt hx have hfx : f x ≠ 0 := hf.map_ne_zero_iff.2 hx apply le_antisymm - · rw [le_div_iff₀ hx', ← NNReal.le_div_iff' hfx, tsub_le_iff_right, ← hf.map_eq_zero, + · rw [le_div_iff₀ hx', ← le_div_iff₀' hfx.bot_lt, tsub_le_iff_right, ← hf.map_eq_zero, hf.map_add, div_mul_cancel₀ _ hfx, hf.map_two, zero_mul] · rw [div_le_iff₀' hx', ← hf.map_eq_zero] refine (mul_eq_zero.1 ?_).resolve_right hfx diff --git a/Mathlib.lean b/Mathlib.lean index 128adb94db271..22ea50c9663f9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -610,6 +610,7 @@ import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Algebra.Order.GroupWithZero.Submonoid import Mathlib.Algebra.Order.GroupWithZero.Synonym import Mathlib.Algebra.Order.GroupWithZero.Unbundled +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 diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean index 34d81f490bea6..73f695ecae794 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean @@ -97,9 +97,9 @@ theorem of_convergence_epsilon : have zero_lt_B : 0 < B := B_ineq.trans_lt' <| mod_cast fib_pos.2 n.succ_pos have nB_pos : 0 < nB := nB_ineq.trans_lt' <| mod_cast fib_pos.2 <| succ_pos _ have zero_lt_mul_conts : 0 < B * nB := by positivity - suffices 1 < ε * (B * nB) from (div_lt_iff zero_lt_mul_conts).mpr this + suffices 1 < ε * (B * nB) from (div_lt_iff₀ zero_lt_mul_conts).mpr this -- use that `N' ≥ n` was obtained from the archimedean property to show the following - calc 1 < ε * (N' : K) := (div_lt_iff' ε_pos).mp one_div_ε_lt_N' + calc 1 < ε * (N' : K) := (div_lt_iff₀' ε_pos).mp one_div_ε_lt_N' _ ≤ ε * (B * nB) := ?_ -- cancel `ε` gcongr diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index 8706c7c0c0d17..c9b87a14420e8 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -223,7 +223,7 @@ variable [LinearOrderedSemifield α] [Archimedean α] {x y ε : α} lemma exists_nat_one_div_lt (hε : 0 < ε) : ∃ n : ℕ, 1 / (n + 1 : α) < ε := by cases' exists_nat_gt (1 / ε) with n hn use n - rw [div_lt_iff, ← div_lt_iff' hε] + rw [div_lt_iff₀, ← div_lt_iff₀' hε] · apply hn.trans simp [zero_lt_one] · exact n.cast_add_one_pos @@ -299,11 +299,11 @@ theorem exists_rat_btwn {x y : α} (h : x < y) : ∃ q : ℚ, x < q ∧ (q : α) refine ⟨(z + 1 : ℤ) / n, ?_⟩ have n0' := (inv_pos.2 (sub_pos.2 h)).trans nh have n0 := Nat.cast_pos.1 n0' - rw [Rat.cast_div_of_ne_zero, Rat.cast_natCast, Rat.cast_intCast, div_lt_iff n0'] - · refine ⟨(lt_div_iff n0').2 <| (lt_iff_lt_of_le_iff_le (zh _)).1 (lt_add_one _), ?_⟩ + rw [Rat.cast_div_of_ne_zero, Rat.cast_natCast, Rat.cast_intCast, div_lt_iff₀ n0'] + · refine ⟨(lt_div_iff₀ n0').2 <| (lt_iff_lt_of_le_iff_le (zh _)).1 (lt_add_one _), ?_⟩ rw [Int.cast_add, Int.cast_one] refine lt_of_le_of_lt (add_le_add_right ((zh _).1 le_rfl) _) ?_ - rwa [← lt_sub_iff_add_lt', ← sub_mul, ← div_lt_iff' (sub_pos.2 h), one_div] + rwa [← lt_sub_iff_add_lt', ← sub_mul, ← div_lt_iff₀' (sub_pos.2 h), one_div] · rw [Rat.den_intCast, Nat.cast_one] exact one_ne_zero · intro H @@ -352,7 +352,7 @@ variable [LinearOrderedField α] theorem archimedean_iff_nat_lt : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x < n := ⟨@exists_nat_gt α _, fun H => ⟨fun x y y0 => - (H (x / y)).imp fun n h => le_of_lt <| by rwa [div_lt_iff y0, ← nsmul_eq_mul] at h⟩⟩ + (H (x / y)).imp fun n h => le_of_lt <| by rwa [div_lt_iff₀ y0, ← nsmul_eq_mul] at h⟩⟩ theorem archimedean_iff_nat_le : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x ≤ n := archimedean_iff_nat_lt.trans diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 2385c27aac032..72a133e1ff7a7 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -5,6 +5,7 @@ Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.Order.Field.Defs +import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Order.Bounds.OrderIso import Mathlib.Tactic.Bound.Attribute @@ -23,55 +24,46 @@ section LinearOrderedSemifield variable [LinearOrderedSemifield α] {a b c d e : α} {m n : ℤ} -/-- `Equiv.mulLeft₀` as an order_iso. -/ -@[simps! (config := { simpRhs := true })] -def OrderIso.mulLeft₀ (a : α) (ha : 0 < a) : α ≃o α := - { Equiv.mulLeft₀ a ha.ne' with map_rel_iff' := @fun _ _ => mul_le_mul_left ha } - -/-- `Equiv.mulRight₀` as an order_iso. -/ -@[simps! (config := { simpRhs := true })] -def OrderIso.mulRight₀ (a : α) (ha : 0 < a) : α ≃o α := - { Equiv.mulRight₀ a ha.ne' with map_rel_iff' := @fun _ _ => mul_le_mul_right ha } - /-! ### Relating one division with another term. -/ -theorem lt_div_iff (hc : 0 < c) : a < b / c ↔ a * c < b := - lt_iff_lt_of_le_iff_le <| div_le_iff₀ hc - -theorem lt_div_iff' (hc : 0 < c) : a < b / c ↔ c * a < b := by rw [mul_comm, lt_div_iff hc] - -theorem div_lt_iff (hc : 0 < c) : b / c < a ↔ b < a * c := - lt_iff_lt_of_le_iff_le (le_div_iff₀ hc) +@[deprecated lt_div_iff₀ (since := "2024-10-02")] +theorem lt_div_iff (hc : 0 < c) : a < b / c ↔ a * c < b := lt_div_iff₀ hc -theorem div_lt_iff' (hc : 0 < c) : b / c < a ↔ b < c * a := by rw [mul_comm, div_lt_iff hc] +@[deprecated lt_div_iff₀' (since := "2024-10-02")] +theorem lt_div_iff' (hc : 0 < c) : a < b / c ↔ c * a < b := lt_div_iff₀' hc -lemma div_lt_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b < c ↔ a / c < b := by - rw [div_lt_iff hb, div_lt_iff' hc] +@[deprecated div_lt_iff₀ (since := "2024-10-02")] +theorem div_lt_iff (hc : 0 < c) : b / c < a ↔ b < a * c := div_lt_iff₀ hc -theorem inv_mul_le_iff (h : 0 < b) : b⁻¹ * a ≤ c ↔ a ≤ b * c := by - rw [inv_eq_one_div, mul_comm, ← div_eq_mul_one_div] - exact div_le_iff₀' h +@[deprecated div_lt_iff₀' (since := "2024-10-02")] +theorem div_lt_iff' (hc : 0 < c) : b / c < a ↔ b < c * a := div_lt_iff₀' hc -theorem inv_mul_le_iff' (h : 0 < b) : b⁻¹ * a ≤ c ↔ a ≤ c * b := by rw [inv_mul_le_iff h, mul_comm] +@[deprecated inv_mul_le_iff₀ (since := "2024-10-02")] +theorem inv_mul_le_iff (h : 0 < b) : b⁻¹ * a ≤ c ↔ a ≤ b * c := inv_mul_le_iff₀ h -theorem mul_inv_le_iff (h : 0 < b) : a * b⁻¹ ≤ c ↔ a ≤ b * c := by rw [mul_comm, inv_mul_le_iff h] +set_option linter.docPrime false in +@[deprecated inv_mul_le_iff₀' (since := "2024-10-02")] +theorem inv_mul_le_iff' (h : 0 < b) : b⁻¹ * a ≤ c ↔ a ≤ c * b := inv_mul_le_iff₀' h -theorem mul_inv_le_iff' (h : 0 < b) : a * b⁻¹ ≤ c ↔ a ≤ c * b := by rw [mul_comm, inv_mul_le_iff' h] +@[deprecated mul_inv_le_iff₀' (since := "2024-10-02")] +theorem mul_inv_le_iff (h : 0 < b) : a * b⁻¹ ≤ c ↔ a ≤ b * c := mul_inv_le_iff₀' h -theorem div_self_le_one (a : α) : a / a ≤ 1 := - if h : a = 0 then by simp [h] else by simp [h] +@[deprecated mul_inv_le_iff₀ (since := "2024-10-02")] +theorem mul_inv_le_iff' (h : 0 < b) : a * b⁻¹ ≤ c ↔ a ≤ c * b := mul_inv_le_iff₀ h -theorem inv_mul_lt_iff (h : 0 < b) : b⁻¹ * a < c ↔ a < b * c := by - rw [inv_eq_one_div, mul_comm, ← div_eq_mul_one_div] - exact div_lt_iff' h +@[deprecated inv_mul_lt_iff₀ (since := "2024-10-02")] +theorem inv_mul_lt_iff (h : 0 < b) : b⁻¹ * a < c ↔ a < b * c := inv_mul_lt_iff₀ h -theorem inv_mul_lt_iff' (h : 0 < b) : b⁻¹ * a < c ↔ a < c * b := by rw [inv_mul_lt_iff h, mul_comm] +@[deprecated inv_mul_lt_iff₀' (since := "2024-10-02")] +theorem inv_mul_lt_iff' (h : 0 < b) : b⁻¹ * a < c ↔ a < c * b := inv_mul_lt_iff₀' h -theorem mul_inv_lt_iff (h : 0 < b) : a * b⁻¹ < c ↔ a < b * c := by rw [mul_comm, inv_mul_lt_iff h] +@[deprecated mul_inv_lt_iff₀' (since := "2024-10-02")] +theorem mul_inv_lt_iff (h : 0 < b) : a * b⁻¹ < c ↔ a < b * c := mul_inv_lt_iff₀' h -theorem mul_inv_lt_iff' (h : 0 < b) : a * b⁻¹ < c ↔ a < c * b := by rw [mul_comm, inv_mul_lt_iff' h] +@[deprecated mul_inv_lt_iff₀ (since := "2024-10-02")] +theorem mul_inv_lt_iff' (h : 0 < b) : a * b⁻¹ < c ↔ a < c * b := mul_inv_lt_iff₀ h theorem inv_pos_le_iff_one_le_mul (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ b * a := by rw [inv_eq_one_div] @@ -83,11 +75,11 @@ theorem inv_pos_le_iff_one_le_mul' (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ a * b : theorem inv_pos_lt_iff_one_lt_mul (ha : 0 < a) : a⁻¹ < b ↔ 1 < b * a := by rw [inv_eq_one_div] - exact div_lt_iff ha + exact div_lt_iff₀ ha theorem inv_pos_lt_iff_one_lt_mul' (ha : 0 < a) : a⁻¹ < b ↔ 1 < a * b := by rw [inv_eq_one_div] - exact div_lt_iff' ha + exact div_lt_iff₀' ha /-- One direction of `div_le_iff` where `b` is allowed to be `0` (but `c` must be nonnegative) -/ theorem div_le_of_nonneg_of_le_mul (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * b) : a / b ≤ c := by @@ -237,7 +229,7 @@ theorem div_le_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c le_iff_le_iff_lt_iff_lt.2 (div_lt_div_left ha hc hb) theorem div_lt_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b < c / d ↔ a * d < c * b := by - rw [lt_div_iff d0, div_mul_eq_mul_div, div_lt_iff b0] + rw [lt_div_iff₀ d0, div_mul_eq_mul_div, div_lt_iff₀ b0] theorem div_le_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := by rw [le_div_iff₀ d0, div_mul_eq_mul_div, div_le_iff₀ b0] @@ -275,9 +267,9 @@ theorem one_le_div (hb : 0 < b) : 1 ≤ a / b ↔ b ≤ a := by rw [le_div_iff theorem div_le_one (hb : 0 < b) : a / b ≤ 1 ↔ a ≤ b := by rw [div_le_iff₀ hb, one_mul] -theorem one_lt_div (hb : 0 < b) : 1 < a / b ↔ b < a := by rw [lt_div_iff hb, one_mul] +theorem one_lt_div (hb : 0 < b) : 1 < a / b ↔ b < a := by rw [lt_div_iff₀ hb, one_mul] -theorem div_lt_one (hb : 0 < b) : a / b < 1 ↔ a < b := by rw [div_lt_iff hb, one_mul] +theorem div_lt_one (hb : 0 < b) : a / b < 1 ↔ a < b := by rw [div_lt_iff₀ hb, one_mul] theorem one_div_le (ha : 0 < a) (hb : 0 < b) : 1 / a ≤ b ↔ 1 / b ≤ a := by simpa using inv_le ha hb @@ -300,7 +292,7 @@ theorem one_div_le_one_div_of_le (ha : 0 < a) (h : a ≤ b) : 1 / b ≤ 1 / a := simpa using inv_le_inv_of_le ha h theorem one_div_lt_one_div_of_lt (ha : 0 < a) (h : a < b) : 1 / b < 1 / a := by - rwa [lt_div_iff' ha, ← div_eq_mul_one_div, div_lt_one (ha.trans h)] + rwa [lt_div_iff₀' ha, ← div_eq_mul_one_div, div_lt_one (ha.trans h)] theorem le_of_one_div_le_one_div (ha : 0 < a) (h : 1 / a ≤ 1 / b) : b ≤ a := le_imp_le_of_lt_imp_lt (one_div_lt_one_div_of_lt ha) h @@ -341,7 +333,7 @@ theorem half_le_self_iff : a / 2 ≤ a ↔ 0 ≤ a := by @[simp] theorem half_lt_self_iff : a / 2 < a ↔ 0 < a := by - rw [div_lt_iff (zero_lt_two' α), mul_two, lt_add_iff_pos_left] + rw [div_lt_iff₀ (zero_lt_two' α), mul_two, lt_add_iff_pos_left] alias ⟨_, half_le_self⟩ := half_le_self_iff @@ -355,9 +347,9 @@ theorem one_half_lt_one : (1 / 2 : α) < 1 := theorem two_inv_lt_one : (2⁻¹ : α) < 1 := (one_div _).symm.trans_lt one_half_lt_one -theorem left_lt_add_div_two : a < (a + b) / 2 ↔ a < b := by simp [lt_div_iff, mul_two] +theorem left_lt_add_div_two : a < (a + b) / 2 ↔ a < b := by simp [lt_div_iff₀, mul_two] -theorem add_div_two_lt_right : (a + b) / 2 < b ↔ a < b := by simp [div_lt_iff, mul_two] +theorem add_div_two_lt_right : (a + b) / 2 < b ↔ a < b := by simp [div_lt_iff₀, mul_two] theorem add_thirds (a : α) : a / 3 + a / 3 + a / 3 = a := by rw [div_add_div_same, div_add_div_same, ← two_mul, ← add_one_mul 2 a, two_add_one_eq_three, @@ -385,12 +377,12 @@ theorem div_mul_le_div_mul_of_div_le_div (h : a / b ≤ c / d) (he : 0 ≤ e) : theorem exists_pos_mul_lt {a : α} (h : 0 < a) (b : α) : ∃ c : α, 0 < c ∧ b * c < a := by have : 0 < a / max (b + 1) 1 := div_pos h (lt_max_iff.2 (Or.inr zero_lt_one)) refine ⟨a / max (b + 1) 1, this, ?_⟩ - rw [← lt_div_iff this, div_div_cancel' h.ne'] + rw [← lt_div_iff₀ this, div_div_cancel' h.ne'] exact lt_max_iff.2 (Or.inl <| lt_add_one _) theorem exists_pos_lt_mul {a : α} (h : 0 < a) (b : α) : ∃ c : α, 0 < c ∧ b < c * a := let ⟨c, hc₀, hc⟩ := exists_pos_mul_lt h b; - ⟨c⁻¹, inv_pos.2 hc₀, by rwa [← div_eq_inv_mul, lt_div_iff hc₀]⟩ + ⟨c⁻¹, inv_pos.2 hc₀, by rwa [← div_eq_inv_mul, lt_div_iff₀ hc₀]⟩ lemma monotone_div_right_of_nonneg (ha : 0 ≤ a) : Monotone (· / a) := fun _b _c hbc ↦ div_le_div_of_nonneg_right hbc ha diff --git a/Mathlib/Algebra/Order/Field/Pointwise.lean b/Mathlib/Algebra/Order/Field/Pointwise.lean index 0b5b49e3c6bb9..01d4453c22523 100644 --- a/Mathlib/Algebra/Order/Field/Pointwise.lean +++ b/Mathlib/Algebra/Order/Field/Pointwise.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Yaël Dillies -/ import Mathlib.Algebra.Group.Pointwise.Set.Basic -import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.SMulWithZero /-! @@ -33,7 +33,7 @@ theorem smul_Ioo : r • Ioo a b = Ioo (r • a) (r • b) := by · exact (mul_lt_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r - refine ⟨⟨(lt_div_iff' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, ?_⟩ + refine ⟨⟨(lt_div_iff₀' hr).mpr a_left, (div_lt_iff₀' hr).mpr a_right⟩, ?_⟩ rw [mul_div_cancel₀ _ (ne_of_gt hr)] theorem smul_Icc : r • Icc a b = Icc (r • a) (r • b) := by @@ -59,7 +59,7 @@ theorem smul_Ico : r • Ico a b = Ico (r • a) (r • b) := by · exact (mul_lt_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r - refine ⟨⟨(le_div_iff₀' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, ?_⟩ + refine ⟨⟨(le_div_iff₀' hr).mpr a_left, (div_lt_iff₀' hr).mpr a_right⟩, ?_⟩ rw [mul_div_cancel₀ _ (ne_of_gt hr)] theorem smul_Ioc : r • Ioc a b = Ioc (r • a) (r • b) := by @@ -72,7 +72,7 @@ theorem smul_Ioc : r • Ioc a b = Ioc (r • a) (r • b) := by · exact (mul_le_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r - refine ⟨⟨(lt_div_iff' hr).mpr a_left, (div_le_iff₀' hr).mpr a_right⟩, ?_⟩ + refine ⟨⟨(lt_div_iff₀' hr).mpr a_left, (div_le_iff₀' hr).mpr a_right⟩, ?_⟩ rw [mul_div_cancel₀ _ (ne_of_gt hr)] theorem smul_Ioi : r • Ioi a = Ioi (r • a) := by @@ -84,7 +84,7 @@ theorem smul_Ioi : r • Ioi a = Ioi (r • a) := by · rintro h use x / r constructor - · exact (lt_div_iff' hr).mpr h + · exact (lt_div_iff₀' hr).mpr h · exact mul_div_cancel₀ _ (ne_of_gt hr) theorem smul_Iio : r • Iio a = Iio (r • a) := by @@ -96,7 +96,7 @@ theorem smul_Iio : r • Iio a = Iio (r • a) := by · rintro h use x / r constructor - · exact (div_lt_iff' hr).mpr h + · exact (div_lt_iff₀' hr).mpr h · exact mul_div_cancel₀ _ (ne_of_gt hr) theorem smul_Ici : r • Ici a = Ici (r • a) := by diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index a3689333018fb..18797cc6170d3 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -484,7 +484,7 @@ theorem floor_div_nat (a : α) (n : ℕ) : ⌊a / n⌋₊ = ⌊a⌋₊ / n := by · exact div_nonneg ha n.cast_nonneg constructor · exact cast_div_le.trans (div_le_div_of_nonneg_right (floor_le ha) n.cast_nonneg) - rw [div_lt_iff, add_mul, one_mul, ← cast_mul, ← cast_add, ← floor_lt ha] + rw [div_lt_iff₀, add_mul, one_mul, ← cast_mul, ← cast_add, ← floor_lt ha] · exact lt_div_mul_add hn · exact cast_pos.2 hn @@ -515,7 +515,7 @@ lemma ceil_lt_mul (hb : 1 < b) (hba : ⌈(b - 1)⁻¹⌉₊ / b < a) : ⌈a⌉ obtain hab | hba := le_total a (b - 1)⁻¹ · calc ⌈a⌉₊ ≤ (⌈(b - 1)⁻¹⌉₊ : α) := by gcongr - _ < b * a := by rwa [← div_lt_iff']; positivity + _ < b * a := by rwa [← div_lt_iff₀']; positivity · rw [← sub_pos] at hb calc ⌈a⌉₊ < a + 1 := ceil_lt_add_one <| hba.trans' <| by positivity @@ -1027,7 +1027,7 @@ theorem sub_floor_div_mul_nonneg (a : k) (hb : 0 < b) : 0 ≤ a - ⌊a / b⌋ * theorem sub_floor_div_mul_lt (a : k) (hb : 0 < b) : a - ⌊a / b⌋ * b < b := sub_lt_iff_lt_add.2 <| by -- Porting note: `← one_add_mul` worked in mathlib3 without the argument - rw [← one_add_mul _ b, ← div_lt_iff hb, add_comm] + rw [← one_add_mul _ b, ← div_lt_iff₀ hb, add_comm] exact lt_floor_add_one _ theorem fract_div_natCast_eq_div_natCast_mod {m n : ℕ} : fract ((m : k) / n) = ↑(m % n) / n := by @@ -1261,7 +1261,7 @@ lemma ceil_div_ceil_inv_sub_one (ha : 1 ≤ a) : ⌈⌈(a - 1)⁻¹⌉ / a⌉ = have : 0 < ⌈(a - 1)⁻¹⌉ := ceil_pos.2 <| by positivity refine le_antisymm (ceil_le.2 <| div_le_self (by positivity) ha.le) <| ?_ rw [le_ceil_iff, sub_lt_comm, div_eq_mul_inv, ← mul_one_sub, - ← lt_div_iff (sub_pos.2 <| inv_lt_one ha)] + ← lt_div_iff₀ (sub_pos.2 <| inv_lt_one ha)] convert ceil_lt_add_one _ using 1 field_simp @@ -1269,7 +1269,7 @@ lemma ceil_lt_mul (hb : 1 < b) (hba : ⌈(b - 1)⁻¹⌉ / b < a) : ⌈a⌉ < b obtain hab | hba := le_total a (b - 1)⁻¹ · calc ⌈a⌉ ≤ (⌈(b - 1)⁻¹⌉ : k) := by gcongr - _ < b * a := by rwa [← div_lt_iff']; positivity + _ < b * a := by rwa [← div_lt_iff₀']; positivity · rw [← sub_pos] at hb calc ⌈a⌉ < a + 1 := ceil_lt_add_one _ @@ -1453,7 +1453,7 @@ section LinearOrderedField variable [LinearOrderedField α] [FloorRing α] theorem round_eq (x : α) : round x = ⌊x + 1 / 2⌋ := by - simp_rw [round, (by simp only [lt_div_iff', two_pos] : 2 * fract x < 1 ↔ fract x < 1 / 2)] + simp_rw [round, (by simp only [lt_div_iff₀', two_pos] : 2 * fract x < 1 ↔ fract x < 1 / 2)] cases' lt_or_le (fract x) (1 / 2) with hx hx · conv_rhs => rw [← fract_add_floor x, add_assoc, add_left_comm, floor_int_add] rw [if_pos hx, self_eq_add_right, floor_eq_iff, cast_zero, zero_add] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index f8df22be3c53e..cabc69bb60181 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -185,12 +185,6 @@ theorem inv_mul_lt_of_lt_mul₀ (h : a < b * c) : b⁻¹ * a < c := by theorem mul_lt_right₀ (c : α) (h : a < b) (hc : c ≠ 0) : a * c < b * c := mul_lt_mul_of_pos_right h (zero_lt_iff.2 hc) -theorem inv_lt_one₀ (ha : a ≠ 0) : a⁻¹ < 1 ↔ 1 < a := - inv_lt_one' (a := Units.mk0 a ha) - -theorem one_lt_inv₀ (ha : a ≠ 0) : 1 < a⁻¹ ↔ a < 1 := - one_lt_inv' (a := Units.mk0 a ha) - theorem inv_lt_inv₀ (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ < b⁻¹ ↔ b < a := show (Units.mk0 a ha)⁻¹ < (Units.mk0 b hb)⁻¹ ↔ Units.mk0 b hb < Units.mk0 a ha from have : CovariantClass αˣ αˣ (· * ·) (· < ·) := diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index ce7ed361fd599..2b57798b2b4fd 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1113,6 +1113,14 @@ end CancelMonoidWithZero section GroupWithZero variable [GroupWithZero G₀] +section Preorder +variable [Preorder G₀] [ZeroLEOneClass G₀] + +/-- See `div_self` for the version with equality when `a ≠ 0`. -/ +lemma div_self_le_one (a : G₀) : a / a ≤ 1 := by obtain rfl | ha := eq_or_ne a 0 <;> simp [*] + +end Preorder + section PartialOrder variable [PartialOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] {a b c d : G₀} @@ -1149,10 +1157,12 @@ lemma zpow_pos_of_pos [PosMulStrictMono G₀] (ha : 0 < a) : ∀ n : ℤ, 0 < a section PosMulMono variable [PosMulMono G₀] +/-- See `le_inv_mul_iff₀'` for a version with multiplication on the other side. -/ lemma le_inv_mul_iff₀ (hc : 0 < c) : a ≤ c⁻¹ * b ↔ c * a ≤ b where mp h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_left h hc.le mpr h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_left h (inv_nonneg.2 hc.le) +/-- See `inv_mul_le_iff₀'` for a version with multiplication on the other side. -/ lemma inv_mul_le_iff₀ (hc : 0 < c) : c⁻¹ * b ≤ a ↔ b ≤ c * a where mp h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_left h hc.le mpr h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_left h (inv_nonneg.2 hc.le) @@ -1168,17 +1178,21 @@ end PosMulMono section MulPosMono variable [MulPosMono G₀] +/-- See `le_mul_inv_iff₀'` for a version with multiplication on the other side. -/ lemma le_mul_inv_iff₀ (hc : 0 < c) : a ≤ b * c⁻¹ ↔ a * c ≤ b where mp h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_right h hc.le mpr h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_right h (inv_nonneg.2 hc.le) +/-- See `mul_inv_le_iff₀'` for a version with multiplication on the other side. -/ lemma mul_inv_le_iff₀ (hc : 0 < c) : b * c⁻¹ ≤ a ↔ b ≤ a * c where mp h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_right h hc.le mpr h := by simpa [hc.ne'] using mul_le_mul_of_nonneg_right h (inv_nonneg.2 hc.le) +/-- See `le_div_iff₀'` for a version with multiplication on the other side. -/ lemma le_div_iff₀ (hc : 0 < c) : a ≤ b / c ↔ a * c ≤ b := by rw [div_eq_mul_inv, le_mul_inv_iff₀ hc] +/-- See `div_le_iff₀'` for a version with multiplication on the other side. -/ lemma div_le_iff₀ (hc : 0 < c) : b / c ≤ a ↔ b ≤ a * c := by rw [div_eq_mul_inv, mul_inv_le_iff₀ hc] @@ -1189,6 +1203,50 @@ lemma div_le_one₀ (hb : 0 < b) : a / b ≤ 1 ↔ a ≤ b := by rw [div_le_iff @[deprecated (since := "2024-08-21")] alias div_le_iff := div_le_iff₀ end MulPosMono + +section PosMulStrictMono +variable [PosMulStrictMono G₀] + +/-- See `lt_inv_mul_iff₀'` for a version with multiplication on the other side. -/ +lemma lt_inv_mul_iff₀ (hc : 0 < c) : a < c⁻¹ * b ↔ c * a < b where + mp h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h hc + mpr h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h (inv_pos.2 hc) + +/-- See `inv_mul_lt_iff₀'` for a version with multiplication on the other side. -/ +lemma inv_mul_lt_iff₀ (hc : 0 < c) : c⁻¹ * b < a ↔ b < c * a where + mp h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h hc + mpr h := by simpa [hc.ne'] using mul_lt_mul_of_pos_left h (inv_pos.2 hc) + +lemma one_lt_inv_mul₀ (ha : 0 < a) : 1 < a⁻¹ * b ↔ a < b := by rw [lt_inv_mul_iff₀ ha, mul_one] +lemma inv_mul_lt_one₀ (ha : 0 < a) : a⁻¹ * b < 1 ↔ b < a := by rw [inv_mul_lt_iff₀ ha, mul_one] + +lemma one_lt_inv₀ (ha : 0 < a) : 1 < a⁻¹ ↔ a < 1 := by simpa using one_lt_inv_mul₀ ha (b := 1) +lemma inv_lt_one₀ (ha : 0 < a) : a⁻¹ < 1 ↔ 1 < a := by simpa using inv_mul_lt_one₀ ha (b := 1) + +end PosMulStrictMono + +section MulPosStrictMono +variable [MulPosStrictMono G₀] + +/-- See `lt_mul_inv_iff₀'` for a version with multiplication on the other side. -/ +lemma lt_mul_inv_iff₀ (hc : 0 < c) : a < b * c⁻¹ ↔ a * c < b where + mp h := by simpa [hc.ne'] using mul_lt_mul_of_pos_right h hc + mpr h := by simpa [hc.ne'] using mul_lt_mul_of_pos_right h (inv_pos.2 hc) + +/-- See `mul_inv_lt_iff₀'` for a version with multiplication on the other side. -/ +lemma mul_inv_lt_iff₀ (hc : 0 < c) : b * c⁻¹ < a ↔ b < a * c where + mp h := by simpa [hc.ne'] using mul_lt_mul_of_pos_right h hc + mpr h := by simpa [hc.ne'] using mul_lt_mul_of_pos_right h (inv_pos.2 hc) + +/-- See `lt_div_iff₀'` for a version with multiplication on the other side. -/ +lemma lt_div_iff₀ (hc : 0 < c) : a < b / c ↔ a * c < b := by + rw [div_eq_mul_inv, lt_mul_inv_iff₀ hc] + +/-- See `div_le_iff₀'` for a version with multiplication on the other side. -/ +lemma div_lt_iff₀ (hc : 0 < c) : b / c < a ↔ b < a * c := by + rw [div_eq_mul_inv, mul_inv_lt_iff₀ hc] + +end MulPosStrictMono end PartialOrder section LinearOrder @@ -1227,22 +1285,95 @@ end CommSemigroupHasZero section CommGroupWithZero variable [CommGroupWithZero G₀] -variable [PartialOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] [MulPosMono G₀] {a b c d : G₀} +variable [PartialOrder G₀] [ZeroLEOneClass G₀] [PosMulReflectLT G₀] + +section PosMulMono +variable [PosMulMono G₀] {a b c d : G₀} + +/-- See `le_inv_mul_iff₀` for a version with multiplication on the other side. -/ +lemma le_inv_mul_iff₀' (hc : 0 < c) : a ≤ c⁻¹ * b ↔ c * a ≤ b := by + rw [le_inv_mul_iff₀ hc, mul_comm] + +/-- See `inv_mul_le_iff₀` for a version with multiplication on the other side. -/ +lemma inv_mul_le_iff₀' (hc : 0 < c) : c⁻¹ * b ≤ a ↔ b ≤ a * c := by + rw [inv_mul_le_iff₀ hc, mul_comm] + +/-- See `le_mul_inv_iff₀` for a version with multiplication on the other side. -/ +lemma le_mul_inv_iff₀' (hc : 0 < c) : a ≤ b * c⁻¹ ↔ c * a ≤ b := by + have := posMulMono_iff_mulPosMono.1 ‹_› + rw [le_mul_inv_iff₀ hc, mul_comm] + +/-- See `mul_inv_le_iff₀` for a version with multiplication on the other side. -/ +lemma mul_inv_le_iff₀' (hc : 0 < c) : b * c⁻¹ ≤ a ↔ b ≤ c * a := by + have := posMulMono_iff_mulPosMono.1 ‹_› + rw [mul_inv_le_iff₀ hc, mul_comm] lemma div_le_div₀ (hb : 0 < b) (hd : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := by + have := posMulMono_iff_mulPosMono.1 ‹_› rw [div_le_iff₀ hb, ← mul_div_right_comm, le_div_iff₀ hd] -lemma le_div_iff₀' (hc : 0 < c) : a ≤ b / c ↔ c * a ≤ b := by rw [le_div_iff₀ hc, mul_comm] -lemma div_le_iff₀' (hc : 0 < c) : b / c ≤ a ↔ b ≤ c * a := by rw [div_le_iff₀ hc, mul_comm] +/-- See `le_div_iff₀` for a version with multiplication on the other side. -/ +lemma le_div_iff₀' (hc : 0 < c) : a ≤ b / c ↔ c * a ≤ b := by + have := posMulMono_iff_mulPosMono.1 ‹_› + rw [le_div_iff₀ hc, mul_comm] + +/-- See `div_le_iff₀` for a version with multiplication on the other side. -/ +lemma div_le_iff₀' (hc : 0 < c) : b / c ≤ a ↔ b ≤ c * a := by + have := posMulMono_iff_mulPosMono.1 ‹_› + rw [div_le_iff₀ hc, mul_comm] lemma le_div_comm₀ (ha : 0 < a) (hc : 0 < c) : a ≤ b / c ↔ c ≤ b / a := by + have := posMulMono_iff_mulPosMono.1 ‹_› rw [le_div_iff₀ ha, le_div_iff₀' hc] lemma div_le_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b ≤ c ↔ a / c ≤ b := by + have := posMulMono_iff_mulPosMono.1 ‹_› rw [div_le_iff₀ hb, div_le_iff₀' hc] @[deprecated (since := "2024-08-21")] alias le_div_iff' := le_div_iff₀' @[deprecated (since := "2024-08-21")] alias div_le_iff' := div_le_iff₀' +end PosMulMono + +section PosMulStrictMono +variable [PosMulStrictMono G₀] {a b c : G₀} + +/-- See `lt_inv_mul_iff₀` for a version with multiplication on the other side. -/ +lemma lt_inv_mul_iff₀' (hc : 0 < c) : a < c⁻¹ * b ↔ a * c < b := by + rw [lt_inv_mul_iff₀ hc, mul_comm] + +/-- See `inv_mul_lt_iff₀` for a version with multiplication on the other side. -/ +lemma inv_mul_lt_iff₀' (hc : 0 < c) : c⁻¹ * b < a ↔ b < a * c := by + rw [inv_mul_lt_iff₀ hc, mul_comm] + +/-- See `lt_mul_inv_iff₀` for a version with multiplication on the other side. -/ +lemma lt_mul_inv_iff₀' (hc : 0 < c) : a < b * c⁻¹ ↔ c * a < b := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [lt_mul_inv_iff₀ hc, mul_comm] + +/-- See `mul_inv_lt_iff₀` for a version with multiplication on the other side. -/ +lemma mul_inv_lt_iff₀' (hc : 0 < c) : b * c⁻¹ < a ↔ b < c * a := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [mul_inv_lt_iff₀ hc, mul_comm] + +/-- See `lt_div_iff₀` for a version with multiplication on the other side. -/ +lemma lt_div_iff₀' (hc : 0 < c) : a < b / c ↔ c * a < b := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [lt_div_iff₀ hc, mul_comm] + +/-- See `div_lt_iff₀` for a version with multiplication on the other side. -/ +lemma div_lt_iff₀' (hc : 0 < c) : b / c < a ↔ b < c * a := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [div_lt_iff₀ hc, mul_comm] + +lemma lt_div_comm₀ (ha : 0 < a) (hc : 0 < c) : a < b / c ↔ c < b / a := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [lt_div_iff₀ ha, lt_div_iff₀' hc] + +lemma div_lt_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b < c ↔ a / c < b := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [div_lt_iff₀ hb, div_lt_iff₀' hc] + +end PosMulStrictMono end CommGroupWithZero diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean new file mode 100644 index 0000000000000..edc0cb47b953d --- /dev/null +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean @@ -0,0 +1,27 @@ +/- +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.Pi.Basic +import Mathlib.Algebra.Order.GroupWithZero.Unbundled +import Mathlib.Algebra.GroupWithZero.Units.Equiv +import Mathlib.Order.Hom.Basic + +/-! +# Multiplication by a positive element as an order isomorphism +-/ + +variable {G₀} [GroupWithZero G₀] [Preorder G₀] [ZeroLEOneClass G₀] {a b c d : G₀} + +/-- `Equiv.mulLeft₀` as an order isomorphism. -/ +@[simps! (config := { simpRhs := true })] +def OrderIso.mulLeft₀ [PosMulMono G₀] [PosMulReflectLE G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where + toEquiv := .mulLeft₀ a ha.ne' + map_rel_iff' := mul_le_mul_left ha + +/-- `Equiv.mulRight₀` as an order isomorphism. -/ +@[simps! (config := { simpRhs := true })] +def OrderIso.mulRight₀ [MulPosMono G₀] [MulPosReflectLE G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where + toEquiv := .mulRight₀ a ha.ne' + map_rel_iff' := mul_le_mul_right ha diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index fda3788c247e8..9017527a7eefa 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -128,13 +128,13 @@ theorem isBigO_iff'' {g : α → E'''} : obtain ⟨c, ⟨hc_pos, hc⟩⟩ := h refine ⟨c⁻¹, ⟨by positivity, ?_⟩⟩ filter_upwards [hc] with x hx - rwa [inv_mul_le_iff (by positivity)] + rwa [inv_mul_le_iff₀ (by positivity)] case mpr => rw [isBigO_iff'] obtain ⟨c, ⟨hc_pos, hc⟩⟩ := h refine ⟨c⁻¹, ⟨by positivity, ?_⟩⟩ filter_upwards [hc] with x hx - rwa [← inv_inv c, inv_mul_le_iff (by positivity)] at hx + rwa [← inv_inv c, inv_mul_le_iff₀ (by positivity)] at hx theorem IsBigO.of_bound (c : ℝ) (h : ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖) : f =O[l] g := isBigO_iff.2 ⟨c, h⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index d664e50c7c219..7a5bdee2eef6e 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -72,7 +72,7 @@ instance CStarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where · simpa only [mem_closedBall_zero_iff, norm_smul, one_mul, norm_star] using (NNReal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) · simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, CStarRing.nnnorm_self_mul_star] - rwa [← NNReal.div_lt_iff (mul_pos ha ha).ne', div_eq_mul_inv, mul_inv, ← mul_assoc] + rwa [← div_lt_iff₀ (mul_pos ha ha), div_eq_mul_inv, mul_inv, ← mul_assoc] section CStarProperty diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index ea6f8405c48bf..67b3db1e4cd80 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -144,7 +144,7 @@ theorem surjOn_closedBall_of_nonlinearRightInverse simp only [dist_le_zero] at this rw [this] have If' : (0 : ℝ) < f'symm.nnnorm := by rw [← inv_pos]; exact (NNReal.coe_nonneg _).trans_lt hc - have Icf' : (c : ℝ) * f'symm.nnnorm < 1 := by rwa [inv_eq_one_div, lt_div_iff If'] at hc + have Icf' : (c : ℝ) * f'symm.nnnorm < 1 := by rwa [inv_eq_one_div, lt_div_iff₀ If'] at hc have Jf' : (f'symm.nnnorm : ℝ) ≠ 0 := ne_of_gt If' have Jcf' : (1 : ℝ) - c * f'symm.nnnorm ≠ 0 := by apply ne_of_gt; linarith /- We have to show that `y` can be written as `f x` for some `x ∈ closedBall b ε`. diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 7459bc396a5f2..703da3c9d5c78 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -901,7 +901,7 @@ theorem Convex.mul_sub_lt_image_sub_of_lt_deriv {D : Set ℝ} (hD : Convex ℝ D obtain ⟨a, a_mem, ha⟩ : ∃ a ∈ Ioo x y, deriv f a = (f y - f x) / (y - x) := exists_deriv_eq_slope f hxy (hf.mono hxyD) (hf'.mono hxyD') have : C < (f y - f x) / (y - x) := ha ▸ hf'_gt _ (hxyD' a_mem) - exact (lt_div_iff (sub_pos.2 hxy)).1 this + exact (lt_div_iff₀ (sub_pos.2 hxy)).1 this /-- Let `f : ℝ → ℝ` be a differentiable function. If `C < f'`, then `f` grows faster than `C * x`, i.e., `C * (y - x) < f y - f x` whenever `x < y`. -/ diff --git a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean index da30835b1f395..df41e2606b7e8 100644 --- a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean +++ b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean @@ -286,7 +286,7 @@ theorem difference_quotients_converge_uniformly refine lt_of_le_of_lt ?_ hqε by_cases hyz' : x = y; · simp [hyz', hqpos.le] have hyz : 0 < ‖y - x‖ := by rw [norm_pos_iff]; intro hy'; exact hyz' (eq_of_sub_eq_zero hy').symm - rw [inv_mul_le_iff hyz, mul_comm, sub_sub_sub_comm] + rw [inv_mul_le_iff₀ hyz, mul_comm, sub_sub_sub_comm] simp only [Pi.zero_apply, dist_zero_left] at e refine Convex.norm_image_sub_le_of_norm_hasFDerivWithin_le @@ -385,7 +385,7 @@ theorem hasFDerivAt_of_tendstoUniformlyOnFilter [NeBot l] by_cases hx : x = n.2; · simp [hx] have hnx : 0 < ‖n.2 - x‖ := by rw [norm_pos_iff]; intro hx'; exact hx (eq_of_sub_eq_zero hx').symm - rw [inv_mul_le_iff hnx, mul_comm] + rw [inv_mul_le_iff₀ hnx, mul_comm] simp only [Function.comp_apply, Prod.map_apply'] rw [norm_sub_rev] exact (f' n.1 x - g' x).le_opNorm (n.2 - x) diff --git a/Mathlib/Analysis/Complex/AbelLimit.lean b/Mathlib/Analysis/Complex/AbelLimit.lean index 84347303b29c2..167aa0434d806 100644 --- a/Mathlib/Analysis/Complex/AbelLimit.lean +++ b/Mathlib/Analysis/Complex/AbelLimit.lean @@ -208,7 +208,7 @@ theorem tendsto_tsum_powerSeries_nhdsWithin_stolzSet gcongr; nth_rw 3 [← mul_one ‖_‖] gcongr; exact pow_le_one₀ (norm_nonneg _) zn.le _ ≤ ‖1 - z‖ * (F + 1) := by gcongr; linarith only - _ < _ := by rwa [norm_sub_rev, lt_div_iff (by positivity)] at zd + _ < _ := by rwa [norm_sub_rev, lt_div_iff₀ (by positivity)] at zd have S₂ : ‖1 - z‖ * ∑ i ∈ Ico B₁ (max B₁ B₂), ‖l - s (i + 1)‖ * ‖z‖ ^ i < ε / 4 := calc _ ≤ ‖1 - z‖ * ∑ i ∈ Ico B₁ (max B₁ B₂), ε / 4 / M * ‖z‖ ^ i := by diff --git a/Mathlib/Analysis/Complex/Hadamard.lean b/Mathlib/Analysis/Complex/Hadamard.lean index 88d6c801d518c..37d20a8d07540 100644 --- a/Mathlib/Analysis/Complex/Hadamard.lean +++ b/Mathlib/Analysis/Complex/Hadamard.lean @@ -189,14 +189,14 @@ lemma F_edge_le_one (f : ℂ → E) (ε : ℝ) (hε : ε > 0) (z : ℂ) rcases hz with hz0 | hz1 -- `z.re = 0` · simp only [hz0, zero_sub, Real.rpow_neg_one, neg_zero, Real.rpow_zero, mul_one, - inv_mul_le_iff (sSupNormIm_eps_pos f hε 0)] + inv_mul_le_iff₀ (sSupNormIm_eps_pos f hε 0)] rw [← hz0] apply le_of_lt (norm_lt_sSupNormIm_eps f ε hε _ _ hB) simp only [verticalClosedStrip, mem_preimage, zero_le_one, left_mem_Icc, hz0] -- `z.re = 1` · rw [mem_singleton_iff] at hz1 simp only [hz1, one_mul, Real.rpow_zero, sub_self, Real.rpow_neg_one, - inv_mul_le_iff (sSupNormIm_eps_pos f hε 1), mul_one] + inv_mul_le_iff₀ (sSupNormIm_eps_pos f hε 1), mul_one] rw [← hz1] apply le_of_lt (norm_lt_sSupNormIm_eps f ε hε _ _ hB) simp only [verticalClosedStrip, mem_preimage, zero_le_one, hz1, right_mem_Icc] @@ -304,7 +304,7 @@ lemma norm_le_interpStrip_of_mem_verticalClosedStrip_eps (ε : ℝ) (hε : ε > ‖f z‖ ≤ ‖((ε + sSupNormIm f 0) ^ (1-z) * (ε + sSupNormIm f 1) ^ z : ℂ)‖ := by simp only [F, abs_invInterpStrip _ _ hε, norm_smul, norm_mul, norm_eq_abs, ← ofReal_add, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε _) _, sub_re, one_re] - rw [← mul_inv_le_iff, ← one_mul (((ε + sSupNormIm f 1) ^ z.re)), ← mul_inv_le_iff', + rw [← mul_inv_le_iff₀', ← one_mul (((ε + sSupNormIm f 1) ^ z.re)), ← mul_inv_le_iff₀, ← Real.rpow_neg_one, ← Real.rpow_neg_one] · simp only [← Real.rpow_mul (le_of_lt (sSupNormIm_eps_pos f hε _)), mul_neg, mul_one, neg_sub, mul_assoc] diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index e172dd9cfe031..a457d182cf842 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -133,7 +133,7 @@ theorem horizontal_strip (hfd : DiffContOnCl ℂ f (im ⁻¹' Ioo a b)) rcases hB with ⟨c, hc, B, hO⟩ obtain ⟨d, ⟨hcd, hd₀⟩, hd⟩ : ∃ d, (c < d ∧ 0 < d) ∧ d < π / 2 / b := by simpa only [max_lt_iff] using exists_between (max_lt hc hπb) - have hb' : d * b < π / 2 := (lt_div_iff hb).1 hd + have hb' : d * b < π / 2 := (lt_div_iff₀ hb).1 hd set aff := (fun w => d * (w - a * I) : ℂ → ℂ) set g := fun (ε : ℝ) (w : ℂ) => exp (ε * (exp (aff w) + exp (-aff w))) /- Since `g ε z → 1` as `ε → 0⁻`, it suffices to prove that `‖g ε z • f z‖ ≤ C` diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index f4c7f5c6d05f3..256120e306349 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -261,7 +261,7 @@ instance : MetricSpace ℍ := have h₀ : 0 < R / im z + 1 := one_pos.trans h₁ refine ⟨log (R / im z + 1), Real.log_pos h₁, ?_⟩ refine fun w hw => (dist_coe_le w z).trans_lt ?_ - rwa [← lt_div_iff' z.im_pos, sub_lt_iff_lt_add, ← Real.lt_log_iff_exp_lt h₀] + rwa [← lt_div_iff₀' z.im_pos, sub_lt_iff_lt_add, ← Real.lt_log_iff_exp_lt h₀] theorem im_pos_of_dist_center_le {z : ℍ} {r : ℝ} {w : ℂ} (h : dist w (center z r) ≤ z.im * Real.sinh r) : 0 < w.im := diff --git a/Mathlib/Analysis/Convex/Body.lean b/Mathlib/Analysis/Convex/Body.lean index ab8c493ff10c3..219c7807cb05d 100644 --- a/Mathlib/Analysis/Convex/Body.lean +++ b/Mathlib/Analysis/Convex/Body.lean @@ -212,7 +212,7 @@ theorem iInter_smul_eq_self [T2Space V] {u : ℕ → ℝ≥0} (K : ConvexBody V) rw [show (1 + u n : ℝ) • y - y = (u n : ℝ) • y by rw [add_smul, one_smul, add_sub_cancel_left], norm_smul, Real.norm_eq_abs] specialize hn n le_rfl - rw [_root_.lt_div_iff' hC_pos, mul_comm, NNReal.coe_zero, sub_zero, Real.norm_eq_abs] at hn + rw [lt_div_iff₀' hC_pos, mul_comm, NNReal.coe_zero, sub_zero, Real.norm_eq_abs] at hn refine lt_of_le_of_lt ?_ hn exact mul_le_mul_of_nonneg_left (hC_bdd _ hyK) (abs_nonneg _) · refine Set.mem_iInter.mpr (fun n => Convex.mem_smul_of_zero_mem K.convex h_zero h ?_) diff --git a/Mathlib/Analysis/Convex/Continuous.lean b/Mathlib/Analysis/Convex/Continuous.lean index 746b4cf9313a6..4b2ebd632b395 100644 --- a/Mathlib/Analysis/Convex/Continuous.lean +++ b/Mathlib/Analysis/Convex/Continuous.lean @@ -139,7 +139,7 @@ lemma ConvexOn.continuousOn_tfae (hC : IsOpen C) (hC' : C.Nonempty) (hf : Convex let w := δ⁻¹ • (z - y) + y have hwyz : δ • w + (1 - δ) • y = z := by simp [w, hδ₀.ne', sub_smul] have hw : dist w x₀ < ε := by - simpa [w, ← hx₀', dist_smul₀, abs_of_nonneg, hδ₀.le, inv_mul_lt_iff', hδ₀] + simpa [w, ← hx₀', dist_smul₀, abs_of_nonneg, hδ₀.le, inv_mul_lt_iff₀', hδ₀] calc f z ≤ max (f w) (f y) := hf.le_max_of_mem_segment (hr hw).2 hy ⟨_, _, hδ₀.le, sub_nonneg.2 hδ₁.le, by simp, hwyz⟩ diff --git a/Mathlib/Analysis/Convex/Deriv.lean b/Mathlib/Analysis/Convex/Deriv.lean index 6954ba6c48f2b..4d71b529373a4 100644 --- a/Mathlib/Analysis/Convex/Deriv.lean +++ b/Mathlib/Analysis/Convex/Deriv.lean @@ -95,7 +95,7 @@ theorem StrictMonoOn.exists_slope_lt_deriv {x y : ℝ} {f : ℝ → ℝ} (hf : C apply ne_of_gt exact hf'_mono ⟨hxw, hwy⟩ ⟨hxw.trans hz.1, hz.2⟩ hz.1 refine ⟨b, ⟨hxw.trans hwb, hby⟩, ?_⟩ - simp only [div_lt_iff, hxy, hxw, hwy, sub_pos] at ha hb ⊢ + simp only [div_lt_iff₀, hxy, hxw, hwy, sub_pos] at ha hb ⊢ have : deriv f a * (w - x) < deriv f b * (w - x) := by apply mul_lt_mul _ le_rfl (sub_pos.2 hxw) _ · exact hf'_mono ⟨hxa, haw.trans hwy⟩ ⟨hxw.trans hwb, hby⟩ (haw.trans hwb) @@ -139,7 +139,7 @@ theorem StrictMonoOn.exists_deriv_lt_slope {x y : ℝ} {f : ℝ → ℝ} (hf : C apply ne_of_gt exact hf'_mono ⟨hxw, hwy⟩ ⟨hxw.trans hz.1, hz.2⟩ hz.1 refine ⟨a, ⟨hxa, haw.trans hwy⟩, ?_⟩ - simp only [lt_div_iff, hxy, hxw, hwy, sub_pos] at ha hb ⊢ + simp only [lt_div_iff₀, hxy, hxw, hwy, sub_pos] at ha hb ⊢ have : deriv f a * (y - w) < deriv f b * (y - w) := by apply mul_lt_mul _ le_rfl (sub_pos.2 hwy) _ · exact hf'_mono ⟨hxa, haw.trans hwy⟩ ⟨hxw.trans hwb, hby⟩ (haw.trans hwb) diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index e68cb3085d3d7..b04bce7ae39f2 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -139,7 +139,7 @@ theorem gauge_le_eq (hs₁ : Convex ℝ s) (hs₀ : (0 : E) ∈ s) (hs₂ : Abso suffices (r⁻¹ * δ) • δ⁻¹ • x ∈ s by rwa [smul_smul, mul_inv_cancel_right₀ δ_pos.ne'] at this rw [mem_smul_set_iff_inv_smul_mem₀ δ_pos.ne'] at hδ refine hs₁.smul_mem_of_zero_mem hs₀ hδ ⟨by positivity, ?_⟩ - rw [inv_mul_le_iff hr', mul_one] + rw [inv_mul_le_iff₀ hr', mul_one] exact hδr.le · have hε' := (lt_add_iff_pos_right a).2 (half_pos hε) exact @@ -369,7 +369,7 @@ theorem gauge_lt_of_mem_smul (x : E) (ε : ℝ) (hε : 0 < ε) (hs₂ : IsOpen s gauge s x < ε := by have : ε⁻¹ • x ∈ s := by rwa [← mem_smul_set_iff_inv_smul_mem₀ hε.ne'] have h_gauge_lt := gauge_lt_one_of_mem_of_isOpen hs₂ this - rwa [gauge_smul_of_nonneg (inv_nonneg.2 hε.le), smul_eq_mul, inv_mul_lt_iff hε, mul_one] + rwa [gauge_smul_of_nonneg (inv_nonneg.2 hε.le), smul_eq_mul, inv_mul_lt_iff₀ hε, mul_one] at h_gauge_lt theorem mem_closure_of_gauge_le_one (hc : Convex ℝ s) (hs₀ : 0 ∈ s) (ha : Absorbent ℝ s) @@ -500,7 +500,7 @@ protected theorem Seminorm.gauge_ball (p : Seminorm ℝ E) : gauge (p.ball 0 1) have hpx₂ : 0 < 2 * p x := mul_pos zero_lt_two hpx refine hp.subset ⟨hpx₂, (2 * p x)⁻¹ • x, ?_, smul_inv_smul₀ hpx₂.ne' _⟩ rw [p.mem_ball_zero, map_smul_eq_mul, Real.norm_eq_abs, abs_of_pos (inv_pos.2 hpx₂), - inv_mul_lt_iff hpx₂, mul_one] + inv_mul_lt_iff₀ hpx₂, mul_one] exact lt_mul_of_one_lt_left hpx one_lt_two refine IsGLB.csInf_eq ⟨fun r => ?_, fun r hr => le_of_forall_pos_le_add fun ε hε => ?_⟩ hp · rintro ⟨hr, y, hy, rfl⟩ @@ -512,7 +512,7 @@ protected theorem Seminorm.gauge_ball (p : Seminorm ℝ E) : gauge (p.ball 0 1) add_pos_of_nonneg_of_pos (apply_nonneg _ _) hε refine hr ⟨hpε, (p x + ε)⁻¹ • x, ?_, smul_inv_smul₀ hpε.ne' _⟩ rw [p.mem_ball_zero, map_smul_eq_mul, Real.norm_eq_abs, abs_of_pos (inv_pos.2 hpε), - inv_mul_lt_iff hpε, mul_one] + inv_mul_lt_iff₀ hpε, mul_one] exact lt_add_of_pos_right _ hε theorem Seminorm.gaugeSeminorm_ball (p : Seminorm ℝ E) : diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index 12b11b68057db..dd75c412a4ffc 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -328,7 +328,7 @@ theorem ae_eq_const_or_norm_integral_lt_of_norm_le_const [StrictConvexSpace ℝ simp [ENNReal.toReal_pos_iff, pos_iff_ne_zero, h₀, measure_lt_top] refine (ae_eq_const_or_norm_average_lt_of_norm_le_const h_le).imp_right fun H => ?_ rwa [average_eq, norm_smul, norm_inv, Real.norm_eq_abs, abs_of_pos hμ, ← div_eq_inv_mul, - div_lt_iff' hμ] at H + div_lt_iff₀' hμ] at H /-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `‖f x‖ ≤ C` a.e. on a set `t` of finite measure, then either this function is a.e. equal to its average value on diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index ea7ca88a2e853..38072902d3e91 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -158,7 +158,7 @@ theorem strictConvexOn_of_slope_strict_mono_adjacent (hs : Convex 𝕜 s) simp_rw [div_eq_iff hxz.ne', ← hab] ring rwa [sub_mul, sub_mul, sub_lt_iff_lt_add', ← add_sub_assoc, lt_sub_iff_add_lt, ← mul_add, - sub_add_sub_cancel, ← lt_div_iff hxz, add_div, mul_div_assoc, mul_div_assoc, mul_comm (f x), + sub_add_sub_cancel, ← lt_div_iff₀ hxz, add_div, mul_div_assoc, mul_div_assoc, mul_comm (f x), mul_comm (f z), ha, hb] at this /-- If for any three points `x < y < z`, the slope of the secant line of `f : 𝕜 → 𝕜` on `[x, y]` is @@ -267,7 +267,7 @@ theorem StrictConvexOn.secant_strict_mono_aux1 (hf : StrictConvexOn 𝕜 s f) {x have hxy' : 0 < y - x := by linarith have hyz' : 0 < z - y := by linarith have hxz' : 0 < z - x := by linarith - rw [← lt_div_iff' hxz'] + rw [← lt_div_iff₀' hxz'] have ha : 0 < (z - y) / (z - x) := by positivity have hb : 0 < (y - x) / (z - x) := by positivity calc diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index 922b18543ffe0..b293f4c6a7eb3 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -40,14 +40,14 @@ theorem strictConvexOn_exp : StrictConvexOn ℝ univ exp := by trans exp y · have h1 : 0 < y - x := by linarith have h2 : x - y < 0 := by linarith - rw [div_lt_iff h1] + rw [div_lt_iff₀ h1] calc exp y - exp x = exp y - exp y * exp (x - y) := by rw [← exp_add]; ring_nf _ = exp y * (1 - exp (x - y)) := by ring _ < exp y * -(x - y) := by gcongr; linarith [add_one_lt_exp h2.ne] _ = exp y * (y - x) := by ring · have h1 : 0 < z - y := by linarith - rw [lt_div_iff h1] + rw [lt_div_iff₀ h1] calc exp y * (z - y) < exp y * (exp (z - y) - 1) := by gcongr _ * ?_ @@ -66,7 +66,7 @@ theorem strictConcaveOn_log_Ioi : StrictConcaveOn ℝ (Ioi 0) log := by have hy : 0 < y := hx.trans hxy trans y⁻¹ · have h : 0 < z - y := by linarith - rw [div_lt_iff h] + rw [div_lt_iff₀ h] have hyz' : 0 < z / y := by positivity have hyz'' : z / y ≠ 1 := by contrapose! h @@ -77,7 +77,7 @@ theorem strictConcaveOn_log_Ioi : StrictConcaveOn ℝ (Ioi 0) log := by _ < z / y - 1 := log_lt_sub_one_of_pos hyz' hyz'' _ = y⁻¹ * (z - y) := by field_simp · have h : 0 < y - x := by linarith - rw [lt_div_iff h] + rw [lt_div_iff₀ h] have hxy' : 0 < x / y := by positivity have hxy'' : x / y ≠ 1 := by contrapose! h @@ -105,12 +105,12 @@ theorem one_add_mul_self_lt_rpow_one_add {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ rw [rpow_def_of_pos hs1, ← exp_log hs2] apply exp_strictMono cases' lt_or_gt_of_ne hs' with hs' hs' - · rw [← div_lt_iff hp', ← div_lt_div_right_of_neg hs'] + · rw [← div_lt_iff₀ hp', ← div_lt_div_right_of_neg hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs2 hs1 hs4 hs3 _ using 1 · rw [add_sub_cancel_left, log_one, sub_zero] · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · apply add_lt_add_left (mul_lt_of_one_lt_left hs' hp) - · rw [← div_lt_iff hp', ← div_lt_div_right hs'] + · rw [← div_lt_iff₀ hp', ← div_lt_div_right hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs1 hs2 hs3 hs4 _ using 1 · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · rw [add_sub_cancel_left, log_one, sub_zero] @@ -144,12 +144,12 @@ theorem rpow_one_add_lt_one_add_mul_self {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ rw [rpow_def_of_pos hs1, ← exp_log hs2] apply exp_strictMono cases' lt_or_gt_of_ne hs' with hs' hs' - · rw [← lt_div_iff hp1, ← div_lt_div_right_of_neg hs'] + · rw [← lt_div_iff₀ hp1, ← div_lt_div_right_of_neg hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs1 hs2 hs3 hs4 _ using 1 · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · rw [add_sub_cancel_left, log_one, sub_zero] · apply add_lt_add_left (lt_mul_of_lt_one_left hs' hp2) - · rw [← lt_div_iff hp1, ← div_lt_div_right hs'] + · rw [← lt_div_iff₀ hp1, ← div_lt_div_right hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs2 hs1 hs4 hs3 _ using 1 · rw [add_sub_cancel_left, log_one, sub_zero] · rw [add_sub_cancel_left, div_div, log_one, sub_zero] @@ -175,20 +175,21 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) have hy' : 0 < y ^ p := rpow_pos_of_pos hy _ trans p * y ^ (p - 1) · have q : 0 < y - x := by rwa [sub_pos] - rw [div_lt_iff q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', ← div_rpow hx hy.le, - sub_lt_comm, ← add_sub_cancel_right (x / y) 1, add_comm, add_sub_assoc, ← div_mul_eq_mul_div, - mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, mul_assoc, ← div_eq_inv_mul, - sub_eq_add_neg, ← mul_neg, ← neg_div, neg_sub, _root_.sub_div, div_self hy.ne'] + rw [div_lt_iff₀ q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', + ← div_rpow hx hy.le, sub_lt_comm, ← add_sub_cancel_right (x / y) 1, add_comm, add_sub_assoc, + ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, + mul_assoc, ← div_eq_inv_mul, sub_eq_add_neg, ← mul_neg, ← neg_div, neg_sub, _root_.sub_div, + div_self hy.ne'] apply one_add_mul_self_lt_rpow_one_add _ _ hp · rw [le_sub_iff_add_le, neg_add_cancel, div_nonneg_iff] exact Or.inl ⟨hx, hy.le⟩ · rw [sub_ne_zero] exact ((div_lt_one hy).mpr hxy).ne · have q : 0 < z - y := by rwa [sub_pos] - rw [lt_div_iff q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', ← div_rpow hz hy.le, - lt_sub_iff_add_lt', ← add_sub_cancel_right (z / y) 1, add_comm _ 1, add_sub_assoc, - ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, - mul_assoc, ← div_eq_inv_mul, _root_.sub_div, div_self hy.ne'] + rw [lt_div_iff₀ q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', + ← div_rpow hz hy.le, lt_sub_iff_add_lt', ← add_sub_cancel_right (z / y) 1, add_comm _ 1, + add_sub_assoc, ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, + rpow_neg_one, mul_assoc, ← div_eq_inv_mul, _root_.sub_div, div_self hy.ne'] apply one_add_mul_self_lt_rpow_one_add _ _ hp · rw [le_sub_iff_add_le, neg_add_cancel, div_nonneg_iff] exact Or.inl ⟨hz, hy.le⟩ diff --git a/Mathlib/Analysis/Convex/StrictConvexSpace.lean b/Mathlib/Analysis/Convex/StrictConvexSpace.lean index 29a52bd553f83..32a1cef9cd3a9 100644 --- a/Mathlib/Analysis/Convex/StrictConvexSpace.lean +++ b/Mathlib/Analysis/Convex/StrictConvexSpace.lean @@ -204,5 +204,5 @@ theorem not_sameRay_iff_abs_lt_norm_sub : ¬SameRay ℝ x y ↔ |‖x‖ - ‖y theorem norm_midpoint_lt_iff (h : ‖x‖ = ‖y‖) : ‖(1 / 2 : ℝ) • (x + y)‖ < ‖x‖ ↔ x ≠ y := by rw [norm_smul, Real.norm_of_nonneg (one_div_nonneg.2 zero_le_two), ← inv_eq_one_div, ← - div_eq_inv_mul, div_lt_iff (zero_lt_two' ℝ), mul_two, ← not_sameRay_iff_of_norm_eq h, + div_eq_inv_mul, div_lt_iff₀ (zero_lt_two' ℝ), mul_two, ← not_sameRay_iff_of_norm_eq h, not_sameRay_iff_norm_add_lt, h] diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 43f9326f818a8..17917e748b4da 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1182,7 +1182,7 @@ instance instZeroAtInftyContinuousMapClass : ZeroAtInftyContinuousMapClass 𝓢( intro ε hε use (SchwartzMap.seminorm ℝ 1 0) f / ε intro x hx - rw [div_lt_iff hε] at hx + rw [div_lt_iff₀ hε] at hx have hxpos : 0 < ‖x‖ := by rw [norm_pos_iff'] intro hxzero @@ -1191,7 +1191,7 @@ instance instZeroAtInftyContinuousMapClass : ZeroAtInftyContinuousMapClass 𝓢( have := norm_pow_mul_le_seminorm ℝ f 1 x rw [pow_one, ← le_div_iff₀' hxpos] at this apply lt_of_le_of_lt this - rwa [div_lt_iff' hxpos] + rwa [div_lt_iff₀' hxpos] /-- Schwartz functions as continuous functions vanishing at infinity. -/ def toZeroAtInfty (f : 𝓢(E, F)) : C₀(E, F) where diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 024a1e6a9ca37..5b128034dec9b 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -140,7 +140,7 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support have : ‖(1 / 2 : ℂ)‖ = 2⁻¹ := by norm_num rw [fourierIntegral_eq_half_sub_half_period_translate hw_ne (hf1.integrable_of_hasCompactSupport hf2), - norm_smul, this, inv_mul_eq_div, div_lt_iff' two_pos] + norm_smul, this, inv_mul_eq_div, div_lt_iff₀' two_pos] refine lt_of_le_of_lt (norm_integral_le_integral_norm _) ?_ simp_rw [Circle.norm_smul] --* Show integral can be taken over A only. @@ -164,8 +164,8 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support simp_rw [norm_norm] simp_rw [dist_eq_norm] at hδ2 refine fun x _ => (hδ2 ?_).le - rw [sub_add_cancel_left, norm_neg, hw'_nm, ← div_div, div_lt_iff (norm_pos_iff.mpr hw_ne), ← - div_lt_iff' hδ1, div_div] + rw [sub_add_cancel_left, norm_neg, hw'_nm, ← div_div, div_lt_iff₀ (norm_pos_iff.mpr hw_ne), ← + div_lt_iff₀' hδ1, div_div] exact (lt_add_of_pos_left _ one_half_pos).trans_le hw_bd have bdA2 := norm_setIntegral_le_of_norm_le_const (hB_vol.trans_lt ENNReal.coe_lt_top) bdA ?_ swap @@ -177,7 +177,7 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support Real.norm_of_nonneg (setIntegral_nonneg mA fun x _ => norm_nonneg _) rw [this] at bdA2 refine bdA2.trans_lt ?_ - rw [div_mul_eq_mul_div, div_lt_iff (NNReal.coe_pos.mpr hB_pos), mul_comm (2 : ℝ), mul_assoc, + rw [div_mul_eq_mul_div, div_lt_iff₀ (NNReal.coe_pos.mpr hB_pos), mul_comm (2 : ℝ), mul_assoc, mul_lt_mul_left hε] rw [← ENNReal.toReal_le_toReal] at hB_vol · refine hB_vol.trans_lt ?_ diff --git a/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean index 7ac440395d2e8..9aa623e738fae 100644 --- a/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean +++ b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean @@ -64,7 +64,7 @@ theorem antilipschitz (coercive : IsCoercive B) : ∃ C : ℝ≥0, 0 < C ∧ Ant refine ⟨C⁻¹.toNNReal, Real.toNNReal_pos.mpr (inv_pos.mpr C_pos), ?_⟩ refine ContinuousLinearMap.antilipschitz_of_bound B♯ ?_ simp_rw [Real.coe_toNNReal', max_eq_left_of_lt (inv_pos.mpr C_pos), ← - inv_mul_le_iff (inv_pos.mpr C_pos)] + inv_mul_le_iff₀ (inv_pos.mpr C_pos)] simpa using below_bound theorem ker_eq_bot (coercive : IsCoercive B) : ker B♯ = ⊥ := by diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index ae6a4a37290bc..b4165a1a77053 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -135,7 +135,7 @@ theorem basisSets_smul_right (v : E) (U : Set E) (hU : U ∈ p.basisSets) : rw [hU, Filter.eventually_iff] simp_rw [(s.sup p).mem_ball_zero, map_smul_eq_mul] by_cases h : 0 < (s.sup p) v - · simp_rw [(lt_div_iff h).symm] + · simp_rw [(lt_div_iff₀ h).symm] rw [← _root_.ball_zero_eq] exact Metric.ball_mem_nhds 0 (div_pos hr h) simp_rw [le_antisymm (not_lt.mp h) (apply_nonneg _ v), mul_zero, hr] diff --git a/Mathlib/Analysis/Normed/Affine/AddTorsor.lean b/Mathlib/Analysis/Normed/Affine/AddTorsor.lean index caa4e080fffdf..2e2257c96f378 100644 --- a/Mathlib/Analysis/Normed/Affine/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Affine/AddTorsor.lean @@ -220,7 +220,7 @@ theorem eventually_homothety_mem_of_mem_interior (x : Q) {s : Set Q} {y : Q} (hy obtain ⟨u, hu₁, hu₂, hu₃⟩ := mem_interior.mp hy obtain ⟨ε, hε, hyε⟩ := Metric.isOpen_iff.mp hu₂ y hu₃ refine ⟨ε / ‖y -ᵥ x‖, div_pos hε hxy, fun δ (hδ : ‖δ - 1‖ < ε / ‖y -ᵥ x‖) => hu₁ (hyε ?_)⟩ - rw [lt_div_iff hxy, ← norm_smul, sub_smul, one_smul] at hδ + rw [lt_div_iff₀ hxy, ← norm_smul, sub_smul, one_smul] at hδ rwa [homothety_apply, Metric.mem_ball, dist_eq_norm_vsub W, vadd_vsub_eq_sub_vsub] theorem eventually_homothety_image_subset_of_finite_subset_interior (x : Q) {s : Set Q} {t : Set Q} diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 61bab9e550cc3..fa49b3701161c 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -101,7 +101,7 @@ theorem mem_resolventSet_of_norm_lt_mul {a : A} {k : 𝕜} (h : ‖a‖ * ‖(1 ne_zero_of_norm_ne_zero ((mul_nonneg (norm_nonneg _) (norm_nonneg _)).trans_lt h).ne' letI ku := Units.map ↑ₐ.toMonoidHom (Units.mk0 k hk) rw [← inv_inv ‖(1 : A)‖, - mul_inv_lt_iff (inv_pos.2 <| norm_pos_iff.2 (one_ne_zero : (1 : A) ≠ 0))] at h + mul_inv_lt_iff₀' (inv_pos.2 <| norm_pos_iff.2 (one_ne_zero : (1 : A) ≠ 0))] at h have hku : ‖-a‖ < ‖(↑ku⁻¹ : A)‖⁻¹ := by simpa [ku, norm_algebraMap] using h simpa [ku, sub_eq_add_neg, Algebra.algebraMap_eq_smul_one] using (ku.add (-a) hku).isUnit diff --git a/Mathlib/Analysis/Normed/Field/ProperSpace.lean b/Mathlib/Analysis/Normed/Field/ProperSpace.lean index 2c8c91cffa8cd..58a51d47a39f0 100644 --- a/Mathlib/Analysis/Normed/Field/ProperSpace.lean +++ b/Mathlib/Analysis/Normed/Field/ProperSpace.lean @@ -42,7 +42,7 @@ lemma ProperSpace.of_nontriviallyNormedField_of_weaklyLocallyCompactSpace ext simp only [mem_closedBall, dist_zero_right, Set.mem_smul_set_iff_inv_smul_mem₀ this, smul_eq_mul, norm_mul, norm_inv, norm_pow, - inv_mul_le_iff (by simpa only [norm_pow] using norm_pos_iff.mpr this)] + inv_mul_le_iff₀ (by simpa only [norm_pow] using norm_pos_iff.mpr this)] have hTop : Tendsto (fun n ↦ ‖c‖^n * r) atTop atTop := Tendsto.atTop_mul_const rpos (tendsto_pow_atTop_atTop_of_one_lt hc) exact .of_seq_closedBall hTop (Eventually.of_forall hC) diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index 0759e6a9c2fc6..9acc022a5fcfb 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -282,9 +282,9 @@ theorem _root_.QuotientAddGroup.norm_lift_apply_le {S : AddSubgroup M} (f : Norm rcases mk_surjective x with ⟨x, rfl⟩ simpa [h] using le_opNorm f x | inr h => - rw [← not_lt, ← _root_.lt_div_iff' h, norm_lt_iff] + rw [← not_lt, ← lt_div_iff₀' h, norm_lt_iff] rintro ⟨x, rfl, hx⟩ - exact ((lt_div_iff' h).1 hx).not_le (le_opNorm f x) + exact ((lt_div_iff₀' h).1 hx).not_le (le_opNorm f x) /-- The operator norm of the projection is `1` if the subspace is not dense. -/ theorem norm_normedMk (S : AddSubgroup M) (h : (S.topologicalClosure : Set M) ≠ univ) : diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index a14d459bc6280..b35a5c4e3baf1 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -186,7 +186,7 @@ theorem NormedSpace.exists_lt_norm (c : ℝ) : ∃ x : E, c < ‖x‖ := by rcases exists_ne (0 : E) with ⟨x, hx⟩ rcases NormedField.exists_lt_norm 𝕜 (c / ‖x‖) with ⟨r, hr⟩ use r • x - rwa [norm_smul, ← _root_.div_lt_iff] + rwa [norm_smul, ← div_lt_iff₀] rwa [norm_pos_iff] protected theorem NormedSpace.unbounded_univ : ¬Bornology.IsBounded (univ : Set E) := fun h => diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean index 7461f70c18dfc..020de596c8f44 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean @@ -156,7 +156,7 @@ theorem exists_lt_apply_of_lt_opNNNorm {𝕜 𝕜₂ E F : Type*} [NormedAddComm obtain ⟨k, hk₁, hk₂⟩ := NormedField.exists_lt_nnnorm_lt 𝕜 hy refine ⟨k • y, (nnnorm_smul k y).symm ▸ (NNReal.lt_inv_iff_mul_lt hy').1 hk₂, ?_⟩ have : ‖σ₁₂ k‖₊ = ‖k‖₊ := Subtype.ext RingHomIsometric.is_iso - rwa [map_smulₛₗ f, nnnorm_smul, ← NNReal.div_lt_iff hfy, div_eq_mul_inv, this] + rwa [map_smulₛₗ f, nnnorm_smul, ← div_lt_iff₀ hfy.bot_lt, div_eq_mul_inv, this] @[deprecated (since := "2024-02-02")] alias exists_lt_apply_of_lt_op_nnnorm := exists_lt_apply_of_lt_opNNNorm diff --git a/Mathlib/Analysis/NormedSpace/Pointwise.lean b/Mathlib/Analysis/NormedSpace/Pointwise.lean index bb061dfe9e186..866ca33561073 100644 --- a/Mathlib/Analysis/NormedSpace/Pointwise.lean +++ b/Mathlib/Analysis/NormedSpace/Pointwise.lean @@ -78,7 +78,7 @@ theorem smul_ball {c : 𝕜} (hc : c ≠ 0) (x : E) (r : ℝ) : c • ball x r = ext y rw [mem_smul_set_iff_inv_smul_mem₀ hc] conv_lhs => rw [← inv_smul_smul₀ hc x] - simp [← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hc), mul_comm _ r, dist_smul₀] + simp [← div_eq_inv_mul, div_lt_iff₀ (norm_pos_iff.2 hc), mul_comm _ r, dist_smul₀] theorem smul_unitBall {c : 𝕜} (hc : c ≠ 0) : c • ball (0 : E) (1 : ℝ) = ball (0 : E) ‖c‖ := by rw [_root_.smul_ball hc, smul_zero, mul_one] diff --git a/Mathlib/Analysis/NormedSpace/RieszLemma.lean b/Mathlib/Analysis/NormedSpace/RieszLemma.lean index 415a55f95a3f4..553d89feaaa42 100644 --- a/Mathlib/Analysis/NormedSpace/RieszLemma.lean +++ b/Mathlib/Analysis/NormedSpace/RieszLemma.lean @@ -50,7 +50,7 @@ theorem riesz_lemma {F : Subspace 𝕜 E} (hFc : IsClosed (F : Set E)) (hF : ∃ simp only [r', max_lt_iff, hr, true_and] norm_num have hlt : 0 < r' := lt_of_lt_of_le (by norm_num) (le_max_right r 2⁻¹) - have hdlt : d < d / r' := (lt_div_iff hlt).mpr ((mul_lt_iff_lt_one_right hdp).2 hr') + have hdlt : d < d / r' := (lt_div_iff₀ hlt).mpr ((mul_lt_iff_lt_one_right hdp).2 hr') obtain ⟨y₀, hy₀F, hxy₀⟩ : ∃ y ∈ F, dist x y < d / r' := (Metric.infDist_lt_iff hFn).mp hdlt have x_ne_y₀ : x - y₀ ∉ F := by by_contra h @@ -63,7 +63,7 @@ theorem riesz_lemma {F : Subspace 𝕜 E} (hFc : IsClosed (F : Set E)) (hF : ∃ r * ‖x - y₀‖ ≤ r' * ‖x - y₀‖ := by gcongr; apply le_max_left _ < d := by rw [← dist_eq_norm] - exact (lt_div_iff' hlt).1 hxy₀ + exact (lt_div_iff₀' hlt).1 hxy₀ _ ≤ dist x (y₀ + y) := Metric.infDist_le_dist_of_mem hy₀y _ = ‖x - y₀ - y‖ := by rw [sub_sub, dist_eq_norm] @@ -82,7 +82,7 @@ theorem riesz_lemma_of_norm_lt {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖ ∃ x₀ : E, ‖x₀‖ ≤ R ∧ ∀ y ∈ F, 1 ≤ ‖x₀ - y‖ := by have Rpos : 0 < R := (norm_nonneg _).trans_lt hR have : ‖c‖ / R < 1 := by - rw [div_lt_iff Rpos] + rw [div_lt_iff₀ Rpos] simpa using hR rcases riesz_lemma hFc hF this with ⟨x, xF, hx⟩ have x0 : x ≠ 0 := fun H => by simp [H] at xF diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index c77e95d8e0c87..c59be01c21689 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -658,7 +658,7 @@ theorem ball_smul (p : Seminorm 𝕜 E) {c : NNReal} (hc : 0 < c) (r : ℝ) (x : (c • p).ball x r = p.ball x (r / c) := by ext rw [mem_ball, mem_ball, smul_apply, NNReal.smul_def, smul_eq_mul, mul_comm, - lt_div_iff (NNReal.coe_pos.mpr hc)] + lt_div_iff₀ (NNReal.coe_pos.mpr hc)] theorem closedBall_smul (p : Seminorm 𝕜 E) {c : NNReal} (hc : 0 < c) (r : ℝ) (x : E) : (c • p).closedBall x r = p.closedBall x (r / c) := by @@ -899,7 +899,7 @@ theorem smul_ball_zero {p : Seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : k ≠ 0) k • p.ball 0 r = p.ball 0 (‖k‖ * r) := by ext rw [mem_smul_set_iff_inv_smul_mem₀ hk, p.mem_ball_zero, p.mem_ball_zero, map_smul_eq_mul, - norm_inv, ← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hk), mul_comm] + norm_inv, ← div_eq_inv_mul, div_lt_iff₀ (norm_pos_iff.2 hk), mul_comm] theorem smul_closedBall_subset {p : Seminorm 𝕜 E} {k : 𝕜} {r : ℝ} : k • p.closedBall 0 r ⊆ p.closedBall 0 (‖k‖ * r) := by @@ -952,7 +952,7 @@ protected theorem absorbent_closedBall (hpr : p x < r) : Absorbent 𝕜 (closedB theorem smul_ball_preimage (p : Seminorm 𝕜 E) (y : E) (r : ℝ) (a : 𝕜) (ha : a ≠ 0) : (a • ·) ⁻¹' p.ball y r = p.ball (a⁻¹ • y) (r / ‖a‖) := Set.ext fun _ => by - rw [mem_preimage, mem_ball, mem_ball, lt_div_iff (norm_pos_iff.mpr ha), mul_comm, ← + rw [mem_preimage, mem_ball, mem_ball, lt_div_iff₀ (norm_pos_iff.mpr ha), mul_comm, ← map_smul_eq_mul p, smul_sub, smul_inv_smul₀ ha] end NormedField @@ -1047,7 +1047,7 @@ theorem continuousAt_zero' [TopologicalSpace E] [ContinuousConstSMul 𝕜 E] {p obtain ⟨k, hk₀, hk⟩ : ∃ k : 𝕜, 0 < ‖k‖ ∧ ‖k‖ * r < ε := by rcases le_or_lt r 0 with hr | hr · use 1; simpa using hr.trans_lt hε - · simpa [lt_div_iff hr] using exists_norm_lt 𝕜 (div_pos hε hr) + · simpa [lt_div_iff₀ hr] using exists_norm_lt 𝕜 (div_pos hε hr) rw [← set_smul_mem_nhds_zero_iff (norm_pos_iff.1 hk₀), smul_closedBall_zero hk₀] at hp exact mem_of_superset hp <| p.closedBall_mono hk.le @@ -1184,9 +1184,9 @@ lemma rescale_to_shell_zpow (p : Seminorm 𝕜 E) {c : 𝕜} (hc : 1 < ‖c‖) refine ⟨-(n+1), ?_, ?_, ?_, ?_⟩ · show c ^ (-(n + 1)) ≠ 0; exact zpow_ne_zero _ (norm_pos_iff.1 cpos) · show p ((c ^ (-(n + 1))) • x) < ε - rw [map_smul_eq_mul, zpow_neg, norm_inv, ← div_eq_inv_mul, div_lt_iff cnpos, mul_comm, + rw [map_smul_eq_mul, zpow_neg, norm_inv, ← div_eq_inv_mul, div_lt_iff₀ cnpos, mul_comm, norm_zpow] - exact (div_lt_iff εpos).1 (hn.2) + exact (div_lt_iff₀ εpos).1 (hn.2) · show ε / ‖c‖ ≤ p (c ^ (-(n + 1)) • x) rw [zpow_neg, div_le_iff₀ cpos, map_smul_eq_mul, norm_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, mul_inv_rev, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel₀ (ne_of_gt cpos), diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean index d9d625ead931a..d5bc0c1f1b5ba 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean @@ -73,7 +73,7 @@ theorem arctan_tan {z : ℂ} (h₀ : z ≠ π / 2) (h₁ : -(π / 2) < z.re) (h rw [← exp_mul_I, ← exp_mul_I, ← exp_sub, show z * I - -z * I = 2 * (I * z) by ring, log_exp, show -I / 2 * (2 * (I * z)) = -(I * I) * z by ring, I_mul_I, neg_neg, one_mul] all_goals norm_num - · rwa [← div_lt_iff' two_pos, neg_div] + · rwa [← div_lt_iff₀' two_pos, neg_div] · rwa [← le_div_iff₀' two_pos] @[simp, norm_cast] diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index d9317023529a2..8d018a967ecb0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -263,7 +263,7 @@ theorem tendsto_exp_div_pow_atTop (n : ℕ) : Tendsto (fun x => exp x / x ^ n) a eventually_atTop.1 ((tendsto_pow_const_div_const_pow_of_one_lt n (one_lt_exp_iff.2 zero_lt_one)).eventually (gt_mem_nhds this)) - simp only [← exp_nat_mul, mul_one, div_lt_iff, exp_pos, ← div_eq_inv_mul] at hN + simp only [← exp_nat_mul, mul_one, div_lt_iff₀, exp_pos, ← div_eq_inv_mul] at hN refine ⟨N, trivial, fun x hx => ?_⟩ rw [Set.mem_Ioi] at hx have hx₀ : 0 < x := (Nat.cast_nonneg N).trans_lt hx diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 79be49beb847d..86a45e8436078 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -85,7 +85,7 @@ theorem finite_integral_rpow_sub_one_pow_aux {r : ℝ} (n : ℕ) (hnr : (n : ℝ refine IntegrableOn.setLIntegral_lt_top ?_ rw [← intervalIntegrable_iff_integrableOn_Ioc_of_le zero_le_one] apply intervalIntegral.intervalIntegrable_rpow' - rwa [neg_lt_neg_iff, inv_mul_lt_iff' hr, one_mul] + rwa [neg_lt_neg_iff, inv_mul_lt_iff₀' hr, one_mul] variable [MeasurableSpace E] [BorelSpace E] {μ : Measure E} [μ.IsAddHaarMeasure] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 23dedfcd33066..f7fafb38f92f6 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -423,7 +423,7 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx intro x hx have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀ refine this ⌊logb r (x / x₀)⌋₊ x ?_ - rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_natCast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add, + rw [mem_Ico, ← div_lt_iff₀ hx₀, ← rpow_natCast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add, Nat.cast_one] exact ⟨hx, Nat.lt_floor_add_one _⟩ intro n diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 705e2df157be4..c5f7769a08f9d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -296,7 +296,7 @@ theorem abs_log_mul_self_lt (x : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) : |log x * x| have : 0 < 1 / x := by simpa only [one_div, inv_pos] using h1 replace := log_le_sub_one_of_pos this replace : log (1 / x) < 1 / x := by linarith - rw [log_div one_ne_zero h1.ne', log_one, zero_sub, lt_div_iff h1] at this + rw [log_div one_ne_zero h1.ne', log_one, zero_sub, lt_div_iff₀ h1] at this have aux : 0 ≤ -log x * x := by refine mul_nonneg ?_ h1.le rw [← log_inv] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 9a2469906a661..c9a10865fd507 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -175,7 +175,7 @@ lemma cpow_ofNat_mul' {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -π < OfNat.ofNa lemma pow_cpow_nat_inv {x : ℂ} {n : ℕ} (h₀ : n ≠ 0) (hlt : -(π / n) < x.arg) (hle : x.arg ≤ π / n) : (x ^ n) ^ (n⁻¹ : ℂ) = x := by rw [← cpow_nat_mul', mul_inv_cancel₀ (Nat.cast_ne_zero.2 h₀), cpow_one] - · rwa [← div_lt_iff' (Nat.cast_pos.2 h₀.bot_lt), neg_div] + · rwa [← div_lt_iff₀' (Nat.cast_pos.2 h₀.bot_lt), neg_div] · rwa [← le_div_iff₀' (Nat.cast_pos.2 h₀.bot_lt)] lemma pow_cpow_ofNat_inv {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -(π / OfNat.ofNat n) < x.arg) diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 0bfe01f772dee..80400c08a6e5d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -823,7 +823,7 @@ theorem rpow_le_one_iff_of_pos (hx : 0 < x) : x ^ y ≤ 1 ↔ 1 ≤ x ∧ y ≤ /-- Bound for `|log x * x ^ t|` in the interval `(0, 1]`, for positive real `t`. -/ theorem abs_log_mul_self_rpow_lt (x t : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) (ht : 0 < t) : |log x * x ^ t| < 1 / t := by - rw [lt_div_iff ht] + rw [lt_div_iff₀ ht] have := abs_log_mul_self_lt (x ^ t) (rpow_pos_of_pos h1 t) (rpow_le_one h1.le h2 ht.le) rwa [log_rpow h1, mul_assoc, abs_mul, abs_of_pos ht, mul_comm] at this diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index 364f5e72076c6..c5ed691712471 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -554,7 +554,7 @@ theorem nsmul_toReal_eq_mul {n : ℕ} (h : n ≠ 0) {θ : Angle} : (n • θ).toReal = n * θ.toReal ↔ θ.toReal ∈ Set.Ioc (-π / n) (π / n) := by nth_rw 1 [← coe_toReal θ] have h' : 0 < (n : ℝ) := mod_cast Nat.pos_of_ne_zero h - rw [← coe_nsmul, nsmul_eq_mul, toReal_coe_eq_self_iff, Set.mem_Ioc, div_lt_iff' h', + rw [← coe_nsmul, nsmul_eq_mul, toReal_coe_eq_self_iff, Set.mem_Ioc, div_lt_iff₀' h', le_div_iff₀' h'] theorem two_nsmul_toReal_eq_two_mul {θ : Angle} : @@ -585,7 +585,7 @@ theorem two_nsmul_toReal_eq_two_mul_sub_two_pi {θ : Angle} : rw [← coe_nsmul, two_nsmul, ← two_mul, toReal_coe_eq_self_sub_two_pi_iff, Set.mem_Ioc] exact ⟨fun h => by linarith, fun h => - ⟨(div_lt_iff' (zero_lt_two' ℝ)).1 h, by linarith [pi_pos, toReal_le_pi θ]⟩⟩ + ⟨(div_lt_iff₀' (zero_lt_two' ℝ)).1 h, by linarith [pi_pos, toReal_le_pi θ]⟩⟩ theorem two_zsmul_toReal_eq_two_mul_sub_two_pi {θ : Angle} : ((2 : ℤ) • θ).toReal = 2 * θ.toReal - 2 * π ↔ π / 2 < θ.toReal := by diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index 2ea81c6afbe95..ee71f3e92b789 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -208,7 +208,7 @@ lemma arctan_add_arctan_lt_pi_div_two {x y : ℝ} (h : x * y < 1) : arctan x + a cases' le_or_lt y 0 with hy hy · rw [← add_zero (π / 2), ← arctan_zero] exact add_lt_add_of_lt_of_le (arctan_lt_pi_div_two _) (tanOrderIso.symm.monotone hy) - · rw [← lt_div_iff hy, ← inv_eq_one_div] at h + · rw [← lt_div_iff₀ hy, ← inv_eq_one_div] at h replace h : arctan x < arctan y⁻¹ := tanOrderIso.symm.strictMono h rwa [arctan_inv_of_pos hy, lt_tsub_iff_right] at h diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index ae6a39293a1df..68cd1103fc6e6 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -236,7 +236,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc simp only [hi.1, mem_Ico, and_true] apply Nat.floor_le_of_le apply le_of_lt - rw [div_lt_iff (Real.log_pos hc), ← Real.log_pow] + rw [div_lt_iff₀ (Real.log_pos hc), ← Real.log_pow] exact Real.log_lt_log hj hi.2 _ = ∑ i ∈ Ico ⌊Real.log j / Real.log c⌋₊ N, (c⁻¹ ^ 2) ^ i := by congr 1 with i diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 9b660a299371d..076f7c450bcde 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -356,7 +356,7 @@ theorem three_le_nValue (hN : 64 ≤ N) : 3 ≤ nValue N := by rw [rpow_natCast] exact (cast_le.2 hN).trans' (by norm_num1) apply lt_of_lt_of_le _ (log_le_log (rpow_pos_of_pos zero_lt_two _) this) - rw [log_rpow zero_lt_two, ← div_lt_iff'] + rw [log_rpow zero_lt_two, ← div_lt_iff₀'] · exact log_two_gt_d9.trans_le' (by norm_num1) · norm_num1 @@ -459,7 +459,7 @@ theorem roth_lower_bound_explicit (hN : 4096 ≤ N) : theorem exp_four_lt : exp 4 < 64 := by rw [show (64 : ℝ) = 2 ^ ((6 : ℕ) : ℝ) by rw [rpow_natCast]; norm_num1, - ← lt_log_iff_exp_lt (rpow_pos_of_pos zero_lt_two _), log_rpow zero_lt_two, ← div_lt_iff'] + ← lt_log_iff_exp_lt (rpow_pos_of_pos zero_lt_two _), log_rpow zero_lt_two, ← div_lt_iff₀'] · exact log_two_gt_d9.trans_le' (by norm_num1) · norm_num diff --git a/Mathlib/Combinatorics/Schnirelmann.lean b/Mathlib/Combinatorics/Schnirelmann.lean index d56f44bd816ff..266093dff23ab 100644 --- a/Mathlib/Combinatorics/Schnirelmann.lean +++ b/Mathlib/Combinatorics/Schnirelmann.lean @@ -196,7 +196,7 @@ lemma schnirelmannDensity_finset (A : Finset ℕ) : schnirelmannDensity A = 0 := let n : ℕ := ⌊A.card / ε⌋₊ + 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] + rw [div_lt_iff₀ (Nat.cast_pos.2 hn), ← div_lt_iff₀' hε, Nat.cast_add_one] exact (Nat.lt_floor_add_one _).trans_le' <| by gcongr; simp [subset_iff] /-- The Schnirelmann density of any finite set is `0`. -/ diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 251ede22462dd..ec59699129256 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -175,8 +175,8 @@ theorem initialBound_pos : 0 < initialBound ε l := theorem hundred_lt_pow_initialBound_mul {ε : ℝ} (hε : 0 < ε) (l : ℕ) : 100 < ↑4 ^ initialBound ε l * ε ^ 5 := by - rw [← rpow_natCast 4, ← div_lt_iff (pow_pos hε 5), lt_rpow_iff_log_lt _ zero_lt_four, ← - div_lt_iff, initialBound, Nat.cast_max, Nat.cast_max] + rw [← rpow_natCast 4, ← div_lt_iff₀ (pow_pos hε 5), lt_rpow_iff_log_lt _ zero_lt_four, ← + div_lt_iff₀, initialBound, Nat.cast_max, Nat.cast_max] · push_cast exact lt_max_of_lt_right (lt_max_of_lt_right <| Nat.lt_floor_add_one _) · exact log_pos (by norm_num) diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean index a6610b54a4ade..572c8e65f679c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean @@ -276,7 +276,7 @@ lemma IsEquipartition.card_interedges_sparsePairs_le' (hP : P.IsEquipartition) · gcongr with UV hUV obtain ⟨U, V⟩ := UV simp [mk_mem_sparsePairs, ← card_interedges_div_card] at hUV - refine ((div_lt_iff ?_).1 hUV.2.2.2).le + refine ((div_lt_iff₀ ?_).1 hUV.2.2.2).le exact mul_pos (Nat.cast_pos.2 (P.nonempty_of_mem_parts hUV.1).card_pos) (Nat.cast_pos.2 (P.nonempty_of_mem_parts hUV.2.1).card_pos) norm_cast diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index cd187d52acfae..9c2622348ebe2 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -28,7 +28,7 @@ theorem isCauSeq_abs_exp (z : ℂ) : let ⟨n, hn⟩ := exists_nat_gt (abs z) have hn0 : (0 : ℝ) < n := lt_of_le_of_lt (abs.nonneg _) hn IsCauSeq.series_ratio_test n (abs z / n) (div_nonneg (abs.nonneg _) (le_of_lt hn0)) - (by rwa [div_lt_iff hn0, one_mul]) fun m hm => by + (by rwa [div_lt_iff₀ hn0, one_mul]) fun m hm => by rw [abs_abs, abs_abs, Nat.factorial_succ, pow_succ', mul_comm m.succ, Nat.cast_mul, ← div_div, mul_div_assoc, mul_div_right_comm, map_mul, map_div₀, abs_natCast] gcongr @@ -1398,7 +1398,7 @@ theorem exp_bound_div_one_sub_of_interval' {x : ℝ} (h1 : 0 < x) (h2 : x < 1) : repeat erw [Finset.sum_range_succ] norm_num [Nat.factorial] nlinarith - _ < 1 / (1 - x) := by rw [lt_div_iff] <;> nlinarith + _ < 1 / (1 - x) := by rw [lt_div_iff₀] <;> nlinarith theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) : Real.exp x ≤ 1 / (1 - x) := by diff --git a/Mathlib/Data/Int/CardIntervalMod.lean b/Mathlib/Data/Int/CardIntervalMod.lean index 045f7ee4c7aa8..b5725e3f43eb6 100644 --- a/Mathlib/Data/Int/CardIntervalMod.lean +++ b/Mathlib/Data/Int/CardIntervalMod.lean @@ -43,14 +43,14 @@ 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 ext x - simp only [mem_map, mem_filter, mem_Ico, ceil_le, lt_ceil, div_le_iff₀, lt_div_iff, + 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 ext x - simp only [mem_map, mem_filter, mem_Ioc, floor_lt, le_floor, div_lt_iff, le_div_iff₀, + 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 @@ -125,7 +125,7 @@ theorem count_modEq_card_eq_ceil (v : ℕ) : rw [← div_add_mod v r, cast_add, cast_mul, add_comm] tactic => simp_rw [← sub_sub, sub_div (_ - _), mul_div_cancel_left₀ _ hr'.ne', ceil_sub_nat] rw [sub_sub_sub_cancel_right, cast_zero, zero_sub] - rw [sub_eq_self, ceil_eq_zero_iff, Set.mem_Ioc, div_le_iff₀ hr', lt_div_iff hr', neg_one_mul, + rw [sub_eq_self, ceil_eq_zero_iff, Set.mem_Ioc, div_le_iff₀ hr', lt_div_iff₀ hr', neg_one_mul, zero_mul, neg_lt_neg_iff, cast_lt] exact ⟨mod_lt _ hr, by simp⟩ @@ -139,10 +139,10 @@ theorem count_modEq_card (v : ℕ) : mul_div_cancel_left₀ _ hr'.ne', add_comm, Int.ceil_add_nat, add_comm] rw [add_right_inj] split_ifs with h - · rw [← cast_sub h.le, Int.ceil_eq_iff, div_le_iff₀ hr', lt_div_iff hr', cast_one, Int.cast_one, + · rw [← cast_sub h.le, Int.ceil_eq_iff, div_le_iff₀ hr', lt_div_iff₀ hr', cast_one, Int.cast_one, sub_self, zero_mul, cast_pos, tsub_pos_iff_lt, one_mul, cast_le, tsub_le_iff_right] exact ⟨h, ((mod_lt _ hr).trans_le (by simp)).le⟩ - · rw [cast_zero, ceil_eq_zero_iff, Set.mem_Ioc, div_le_iff₀ hr', lt_div_iff hr', zero_mul, + · rw [cast_zero, ceil_eq_zero_iff, Set.mem_Ioc, div_le_iff₀ hr', lt_div_iff₀ hr', zero_mul, tsub_nonpos, ← neg_eq_neg_one_mul, neg_lt_sub_iff_lt_add, ← cast_add, cast_lt, cast_le] exact ⟨(mod_lt _ hr).trans_le (by simp), not_lt.mp h⟩ diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index da7622c261a6f..796985dd715ad 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -841,25 +841,25 @@ theorem div_le_of_le_mul' {a b c : ℝ≥0} (h : a ≤ b * c) : a / b ≤ c := @[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₀ <| pos_iff_ne_zero.2 hr + le_div_iff₀ hr.bot_lt -nonrec theorem le_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ r * a ≤ b := - le_div_iff₀' <| pos_iff_ne_zero.2 hr +@[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 -theorem div_lt_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < b * r := - lt_iff_lt_of_le_iff_le (le_div_iff₀ (pos_iff_ne_zero.2 hr)) +@[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 -theorem div_lt_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r < b ↔ a < r * b := - lt_iff_lt_of_le_iff_le (le_div_iff₀' (pos_iff_ne_zero.2 hr)) +@[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 -theorem lt_div_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ a * r < b := - lt_iff_lt_of_le_iff_le (div_le_iff₀ (pos_iff_ne_zero.2 hr)) +@[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 -theorem lt_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ r * a < b := - lt_iff_lt_of_le_iff_le (div_le_iff₀' (pos_iff_ne_zero.2 hr)) +@[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 fun hr => False.elim <| by simp [hr] at h).1 h + (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 := @@ -884,8 +884,7 @@ 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, one_mul] - exact ne_of_gt (lt_of_le_of_lt (zero_le _) h) + 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 @@ -902,7 +901,7 @@ theorem _root_.Real.toNNReal_div' {x y : ℝ} (hy : 0 ≤ y) : 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, one_mul] + rw [← one_div, div_lt_iff₀ hx.bot_lt, one_mul] theorem zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := zpow_pos_of_pos hx.bot_lt _ diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index aab8b86739170..f27de94d155e4 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -68,7 +68,7 @@ theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x : intro n n0 y yS have := (Int.sub_one_lt_floor _).trans_le (Int.cast_le.2 <| (hf n).2 _ ⟨y, yS, Int.floor_le _⟩) simp only [Rat.cast_div, Rat.cast_intCast, Rat.cast_natCast, gt_iff_lt] - rwa [lt_div_iff (Nat.cast_pos.2 n0 : (_ : ℝ) < _), sub_mul, inv_mul_cancel₀] + rwa [lt_div_iff₀ (Nat.cast_pos.2 n0 : (_ : ℝ) < _), sub_mul, inv_mul_cancel₀] exact ne_of_gt (Nat.cast_pos.2 n0) have hg : IsCauSeq abs (fun n => f n / n : ℕ → ℚ) := by intro ε ε0 diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 53573cbcf7b0b..9a77915281ee6 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -23,7 +23,7 @@ namespace Real theorem pi_gt_sqrtTwoAddSeries (n : ℕ) : 2 ^ (n + 1) * √(2 - sqrtTwoAddSeries 0 n) < π := by have : √(2 - sqrtTwoAddSeries 0 n) / 2 * 2 ^ (n + 2) < π := by - rw [← lt_div_iff, ← sin_pi_over_two_pow_succ] + rw [← lt_div_iff₀, ← sin_pi_over_two_pow_succ] focus apply sin_lt apply div_pos pi_pos @@ -34,7 +34,7 @@ theorem pi_gt_sqrtTwoAddSeries (n : ℕ) : 2 ^ (n + 1) * √(2 - sqrtTwoAddSerie theorem pi_lt_sqrtTwoAddSeries (n : ℕ) : π < 2 ^ (n + 1) * √(2 - sqrtTwoAddSeries 0 n) + 1 / 4 ^ n := by have : π < (√(2 - sqrtTwoAddSeries 0 n) / 2 + 1 / (2 ^ n) ^ 3 / 4) * (2 : ℝ) ^ (n + 2) := by - rw [← div_lt_iff (by norm_num), ← sin_pi_over_two_pow_succ] + rw [← div_lt_iff₀ (by norm_num), ← sin_pi_over_two_pow_succ] refine lt_of_lt_of_le (lt_add_of_sub_right_lt (sin_gt_sub_cube ?_ ?_)) ?_ · apply div_pos pi_pos; apply pow_pos; norm_num · rw [div_le_iff₀'] diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index 479d737a940f8..778010911a87f 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -279,13 +279,13 @@ private lemma not_irrational_exists_rep {x : ℝ} : obtain ⟨a, b, hb, h⟩ := not_irrational_exists_rep h' have ha : (0 : ℝ) < a := by have : 0 < (a : ℝ) / b := h ▸ pi_div_two_pos - rwa [lt_div_iff (by positivity), zero_mul] at this + rwa [lt_div_iff₀ (by positivity), zero_mul] at this have k (n : ℕ) : 0 < (a : ℝ) ^ (2 * n + 1) / n ! := by positivity have j : ∀ᶠ n : ℕ in atTop, (a : ℝ) ^ (2 * n + 1) / n ! * I n (π / 2) < 1 := by have := eventually_lt_of_tendsto_lt (show (0 : ℝ) < 1 / 2 by norm_num) (tendsto_pow_div_factorial_at_top_aux a) filter_upwards [this] with n hn - rw [lt_div_iff (zero_lt_two : (0 : ℝ) < 2)] at hn + rw [lt_div_iff₀ (zero_lt_two : (0 : ℝ) < 2)] at hn exact hn.trans_le' (mul_le_mul_of_nonneg_left (I_le _) (by positivity)) obtain ⟨n, hn⟩ := j.exists have hn' : 0 < a ^ (2 * n + 1) / n ! * I n (π / 2) := mul_pos (k _) I_pos diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 1b83b2cd5269f..7ee860b81b0d1 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -509,12 +509,12 @@ variable [LinearOrderedField α] {a : α} @[simp] theorem preimage_mul_const_Iio (a : α) {c : α} (h : 0 < c) : (fun x => x * c) ⁻¹' Iio a = Iio (a / c) := - ext fun _x => (lt_div_iff h).symm + ext fun _x => (lt_div_iff₀ h).symm @[simp] theorem preimage_mul_const_Ioi (a : α) {c : α} (h : 0 < c) : (fun x => x * c) ⁻¹' Ioi a = Ioi (a / c) := - ext fun _x => (div_lt_iff h).symm + ext fun _x => (div_lt_iff₀ h).symm @[simp] theorem preimage_mul_const_Iic (a : α) {c : α} (h : 0 < c) : @@ -582,11 +582,11 @@ theorem preimage_mul_const_Icc_of_neg (a b : α) {c : α} (h : c < 0) : @[simp] theorem preimage_const_mul_Iio (a : α) {c : α} (h : 0 < c) : (c * ·) ⁻¹' Iio a = Iio (a / c) := - ext fun _x => (lt_div_iff' h).symm + ext fun _x => (lt_div_iff₀' h).symm @[simp] theorem preimage_const_mul_Ioi (a : α) {c : α} (h : 0 < c) : (c * ·) ⁻¹' Ioi a = Ioi (a / c) := - ext fun _x => (div_lt_iff' h).symm + ext fun _x => (div_lt_iff₀' h).symm @[simp] theorem preimage_const_mul_Iic (a : α) {c : α} (h : 0 < c) : (c * ·) ⁻¹' Iic a = Iic (a / c) := diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 14fe8b165b42a..6b82e8495e419 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -185,7 +185,7 @@ theorem stereoInvFun_ne_north_pole (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : rw [← inner_lt_one_iff_real_of_norm_one _ hv] · have hw : ⟪v, w⟫_ℝ = 0 := Submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2 have hw' : (‖(w : E)‖ ^ 2 + 4)⁻¹ * (‖(w : E)‖ ^ 2 - 4) < 1 := by - rw [inv_mul_lt_iff'] + rw [inv_mul_lt_iff₀'] · linarith positivity simpa [real_inner_comm, inner_add_right, inner_smul_right, real_inner_self_eq_norm_mul_norm, hw, diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 01130e726c4b7..064e40fa40d71 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -266,7 +266,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a convert h.comp_mul a ext t rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, Real.dist_eq, Real.dist_eq, - lt_div_iff (abs_pos.mpr ha), ← abs_mul, sub_mul, div_mul_cancel₀ _ ha] + lt_div_iff₀ (abs_pos.mpr ha), ← abs_mul, sub_mul, div_mul_cancel₀ _ ha] lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by diff --git a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean index e6dd7b5af91ea..06f6a65c35c43 100644 --- a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean +++ b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean @@ -37,7 +37,7 @@ theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toL refine (h_i ▸ Finset.le_sup' (fun i => ‖v i‖) (Finset.mem_univ j)).trans ?_ exact norm_le_zero_iff.mpr h_nz have h_le : ∀ j, ‖v j * (v i)⁻¹‖ ≤ 1 := fun j => by - rw [norm_mul, norm_inv, mul_inv_le_iff' (norm_pos_iff.mpr h_nz), one_mul] + rw [norm_mul, norm_inv, mul_inv_le_iff₀ (norm_pos_iff.mpr h_nz), one_mul] exact h_i ▸ Finset.le_sup' (fun i => ‖v i‖) (Finset.mem_univ j) simp_rw [mem_closedBall_iff_norm'] refine ⟨i, ?_⟩ diff --git a/Mathlib/MeasureTheory/Covering/DensityTheorem.lean b/Mathlib/MeasureTheory/Covering/DensityTheorem.lean index 51f94e3d8e12d..62c58768b78a6 100644 --- a/Mathlib/MeasureTheory/Covering/DensityTheorem.lean +++ b/Mathlib/MeasureTheory/Covering/DensityTheorem.lean @@ -123,7 +123,7 @@ theorem tendsto_closedBall_filterAt {K : ℝ} {x : α} {ι : Type*} {l : Filter apply (((Metric.tendsto_nhds.mp δlim _ (div_pos hε hK)).and δpos).and xmem).mono rintro j ⟨⟨hjε, hj₀ : 0 < δ j⟩, hx⟩ y hy replace hjε : (K + 1) * δ j < ε := by - simpa [abs_eq_self.mpr hj₀.le] using (lt_div_iff' hK).mp hjε + simpa [abs_eq_self.mpr hj₀.le] using (lt_div_iff₀' hK).mp hjε simp only [mem_closedBall] at hx hy ⊢ linarith [dist_triangle_right y x (w j)] diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 5928bb689a9c9..364031a842a49 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -112,7 +112,7 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S · refine ⟨a, ⟨hat, a_disj⟩, ?_⟩ simpa only [← mzero, zero_div] using δnonneg a hat · have I : m / τ < m := by - rw [div_lt_iff (zero_lt_one.trans hτ)] + rw [div_lt_iff₀ (zero_lt_one.trans hτ)] conv_lhs => rw [← mul_one m] exact (mul_lt_mul_left mpos).2 hτ rcases exists_lt_of_lt_csSup (Anonempty.image _) I with ⟨x, xA, hx⟩ diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index 31b0277e42b2f..2d82b8b7e4d4e 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -203,7 +203,7 @@ theorem TendstoInMeasure.exists_seq_tendsto_ae (hfg : TendstoInMeasure μ f atTo refine ⟨max N (k - 1), fun n hn_ge => lt_of_le_of_lt ?_ hk_lt_ε⟩ specialize hNx n ((le_max_left _ _).trans hn_ge) have h_inv_n_le_k : (2 : ℝ)⁻¹ ^ n ≤ 2 * (2 : ℝ)⁻¹ ^ k := by - rw [mul_comm, ← inv_mul_le_iff' (zero_lt_two' ℝ)] + rw [mul_comm, ← inv_mul_le_iff₀' (zero_lt_two' ℝ)] conv_lhs => congr rw [← pow_one (2 : ℝ)⁻¹] diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index 39a15c80e968f..7b324585c63a5 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -312,7 +312,7 @@ theorem lintegral_Lp_mul_le_Lq_mul_Lr {α} [MeasurableSpace α] {p q r : ℝ} (h let p2 := q / p let q2 := p2.conjExponent have hp2q2 : p2.IsConjExponent q2 := - .conjExponent (by simp [p2, q2, _root_.lt_div_iff, hpq, hp0_lt]) + .conjExponent (by simp [p2, q2, _root_.lt_div_iff₀, hpq, hp0_lt]) calc (∫⁻ a : α, (f * g) a ^ p ∂μ) ^ (1 / p) = (∫⁻ a : α, f a ^ p * g a ^ p ∂μ) ^ (1 / p) := by simp_rw [Pi.mul_apply, ENNReal.mul_rpow_of_nonneg _ _ hp0] diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 8db2aa64138ec..df652367ebf11 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -438,7 +438,7 @@ theorem tendsto_integral_comp_smul_smul_of_integrable simp [norm_smul, abs_of_pos cpos, mul_pow]; ring _ < δ ^ finrank ℝ F * ε := by apply hM - rw [div_lt_iff δpos] at hc + rw [div_lt_iff₀ δpos] at hc simp only [mem_compl_iff, mem_closedBall, dist_zero_right, norm_smul, Real.norm_eq_abs, abs_of_nonneg cpos.le, not_le, gt_iff_lt] exact hc.trans_le (by gcongr) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index a43d6a7133501..110086e7b6708 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -212,7 +212,7 @@ theorem MeasureTheory.volume_sum_rpow_lt [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) simp_rw [← Set.preimage_smul_inv₀ (ne_of_gt hr), Set.preimage_setOf_eq, Pi.smul_apply, smul_eq_mul, abs_mul, mul_rpow (abs_nonneg _) (abs_nonneg _), abs_inv, inv_rpow (abs_nonneg _), ← Finset.mul_sum, abs_eq_self.mpr (le_of_lt hr), - inv_mul_lt_iff (rpow_pos_of_pos hr _), mul_one, ← rpow_lt_rpow_iff + inv_mul_lt_iff₀ (rpow_pos_of_pos hr _), mul_one, ← rpow_lt_rpow_iff (rpow_nonneg (h₁ _) _) (le_of_lt hr) (by linarith : 0 < p), ← rpow_mul (h₁ _), div_mul_cancel₀ _ (ne_of_gt (by linarith) : p ≠ 0), Real.rpow_one] @@ -284,7 +284,7 @@ theorem Complex.volume_sum_rpow_lt [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) (r : convert addHaar_smul_of_nonneg volume (le_of_lt hr) {x : ι → ℂ | ∑ i, ‖x i‖ ^ p < 1} using 2 · simp_rw [← Set.preimage_smul_inv₀ (ne_of_gt hr), Set.preimage_setOf_eq, Pi.smul_apply, norm_smul, mul_rpow (norm_nonneg _) (norm_nonneg _), Real.norm_eq_abs, abs_inv, inv_rpow - (abs_nonneg _), ← Finset.mul_sum, abs_eq_self.mpr (le_of_lt hr), inv_mul_lt_iff + (abs_nonneg _), ← Finset.mul_sum, abs_eq_self.mpr (le_of_lt hr), inv_mul_lt_iff₀ (rpow_pos_of_pos hr _), mul_one, ← rpow_lt_rpow_iff (rpow_nonneg (h₁ _) _) (le_of_lt hr) (by linarith : 0 < p), ← rpow_mul (h₁ _), div_mul_cancel₀ _ (ne_of_gt (by linarith) : p ≠ 0), Real.rpow_one] diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbs.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbs.lean index dd9d6a2408815..694594f1a0e40 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbs.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbs.lean @@ -38,7 +38,7 @@ theorem exists_partition_int (n : ℕ) {ε : ℝ} (hε : 0 < ε) {b : ℤ} (hb : refine ⟨fun i ↦ ⟨natAbs (floor ((A i % b : ℤ) / abs b • ε : ℝ)), ?_⟩, ?_⟩ · rw [← ofNat_lt, natAbs_of_nonneg (hfloor i), floor_lt] apply lt_of_lt_of_le _ (Nat.le_ceil _) - rw [Algebra.smul_def, eq_intCast, ← div_div, div_lt_div_right hε, div_lt_iff hb', one_mul, + rw [Algebra.smul_def, eq_intCast, ← div_div, div_lt_div_right hε, div_lt_iff₀ hb', one_mul, cast_lt] exact Int.emod_lt _ hb intro i₀ i₁ hi @@ -46,7 +46,7 @@ theorem exists_partition_int (n : ℕ) {ε : ℝ} (hε : 0 < ε) {b : ℤ} (hb : congr_arg ((↑) : ℕ → ℤ) (Fin.mk_eq_mk.mp hi) rw [natAbs_of_nonneg (hfloor i₀), natAbs_of_nonneg (hfloor i₁)] at hi have hi := abs_sub_lt_one_of_floor_eq_floor hi - rw [abs_sub_comm, ← sub_div, abs_div, abs_of_nonneg hbε.le, div_lt_iff hbε, one_mul] at hi + rw [abs_sub_comm, ← sub_div, abs_div, abs_of_nonneg hbε.le, div_lt_iff₀ hbε, one_mul] at hi rwa [Int.cast_abs, Int.cast_sub] /-- `abs : ℤ → ℤ` is an admissible absolute value. -/ diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index 7b6a30b0c651c..c1fd34862bfb8 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -131,7 +131,7 @@ theorem exists_approx_polynomial {b : Fq[X]} (hb : b ≠ 0) {ε : ℝ} (hε : 0 cardPowDegree_nonzero _ h', cardPowDegree_nonzero _ hb, Algebra.smul_def, eq_intCast, Int.cast_pow, Int.cast_natCast, Int.cast_pow, Int.cast_natCast, log_mul (pow_ne_zero _ q_pos'.ne') hε.ne', ← rpow_natCast, ← rpow_natCast, log_rpow q_pos', - log_rpow q_pos', ← lt_div_iff (log_pos one_lt_q'), add_div, + log_rpow q_pos', ← lt_div_iff₀ (log_pos one_lt_q'), add_div, mul_div_cancel_right₀ _ (log_pos one_lt_q').ne'] -- And that result follows from manipulating the result from `exists_approx_polynomial_aux` -- to turn the `-⌈-stuff⌉₊` into `+ stuff`. diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 77d109b645420..946ac3af1f703 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -235,7 +235,7 @@ theorem den_le_and_le_num_le_of_sub_lt_one_div_den_sq {ξ q : ℚ} · exact le_rfl · have hξ₀ : (0 : ℚ) < ξ.den := Nat.cast_pos.mpr ξ.pos rw [← Rat.num_div_den ξ, div_mul_eq_mul_div, div_sub' _ _ _ hξ₀.ne', abs_div, abs_of_pos hξ₀, - div_lt_iff hξ₀, div_mul_comm, mul_one] at h + div_lt_iff₀ hξ₀, div_mul_comm, mul_one] at h refine Nat.cast_le.mp ((one_lt_div hq₀).mp <| lt_of_le_of_lt ?_ h).le norm_cast rw [mul_comm _ q.num] @@ -422,9 +422,9 @@ private theorem aux₂ : 0 < u - ⌊ξ⌋ * v ∧ u - ⌊ξ⌋ * v < v := by obtain ⟨hcop, _, h⟩ := h obtain ⟨hv₀, hv₀'⟩ := aux₀ (zero_lt_two.trans_le hv) have hv₁ : 0 < 2 * v - 1 := by linarith only [hv] - rw [← one_div, lt_div_iff (mul_pos hv₀ hv₀'), ← abs_of_pos (mul_pos hv₀ hv₀'), ← abs_mul, sub_mul, - ← mul_assoc, ← mul_assoc, div_mul_cancel₀ _ hv₀.ne', abs_sub_comm, abs_lt, lt_sub_iff_add_lt, - sub_lt_iff_lt_add, mul_assoc] at h + rw [← one_div, lt_div_iff₀ (mul_pos hv₀ hv₀'), ← abs_of_pos (mul_pos hv₀ hv₀'), ← abs_mul, + sub_mul, ← mul_assoc, ← mul_assoc, div_mul_cancel₀ _ hv₀.ne', abs_sub_comm, abs_lt, + lt_sub_iff_add_lt, sub_lt_iff_lt_add, mul_assoc] at h have hu₀ : 0 ≤ u - ⌊ξ⌋ * v := by -- Porting note: this abused the definitional equality `-1 + 1 = 0` -- refine' (mul_nonneg_iff_of_pos_right hv₁).mp ((lt_iff_add_one_le (-1 : ℤ) _).mp _) diff --git a/Mathlib/NumberTheory/Liouville/Basic.lean b/Mathlib/NumberTheory/Liouville/Basic.lean index e76ed12616deb..04403810047a7 100644 --- a/Mathlib/NumberTheory/Liouville/Basic.lean +++ b/Mathlib/NumberTheory/Liouville/Basic.lean @@ -191,11 +191,11 @@ protected theorem transcendental {x : ℝ} (lx : Liouville x) : Transcendental -- recall, this is a proof by contradiction! refine lt_irrefl ((b : ℝ) ^ f.natDegree * |x - ↑a / ↑b|) ?_ -- clear denominators at `a1` - rw [lt_div_iff' (pow_pos b0 _), pow_add, mul_assoc] at a1 + rw [lt_div_iff₀' (pow_pos b0 _), pow_add, mul_assoc] at a1 -- split the inequality via `1 / A`. refine (?_ : (b : ℝ) ^ f.natDegree * |x - a / b| < 1 / A).trans_le ?_ -- This branch of the proof uses the Liouville condition and the Archimedean property - · refine (lt_div_iff' hA).mpr ?_ + · refine (lt_div_iff₀' hA).mpr ?_ refine lt_of_le_of_lt ?_ a1 gcongr refine hn.le.trans ?_ diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean index 45980956c089e..26c83fc23c3a5 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean @@ -53,7 +53,7 @@ theorem liouvilleWith_one (x : ℝ) : LiouvilleWith 1 x := by refine ((eventually_gt_atTop 0).mono fun n hn => ?_).frequently have hn' : (0 : ℝ) < n := by simpa have : x < ↑(⌊x * ↑n⌋ + 1) / ↑n := by - rw [lt_div_iff hn', Int.cast_add, Int.cast_one] + rw [lt_div_iff₀ hn', Int.cast_add, Int.cast_one] exact Int.lt_floor_add_one _ refine ⟨⌊x * n⌋ + 1, this.ne, ?_⟩ rw [abs_sub_comm, abs_of_pos (sub_pos.2 this), rpow_one, sub_lt_iff_lt_add', @@ -100,7 +100,7 @@ theorem frequently_lt_rpow_neg (h : LiouvilleWith p x) (hlt : q < p) : refine (this.and_frequently hC).mono ?_ rintro n ⟨hnC, hn, m, hne, hlt⟩ replace hn : (0 : ℝ) < n := Nat.cast_pos.2 hn - refine ⟨m, hne, hlt.trans <| (div_lt_iff <| rpow_pos_of_pos hn _).2 ?_⟩ + refine ⟨m, hne, hlt.trans <| (div_lt_iff₀ <| rpow_pos_of_pos hn _).2 ?_⟩ rwa [mul_comm, ← rpow_add hn, ← sub_eq_add_neg] /-- The product of a Liouville number and a nonzero rational number is again a Liouville number. -/ diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 812b9760b6b67..0bc6a211e868a 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -357,7 +357,7 @@ theorem normSq_S_smul_lt_one (h : 1 < normSq z) : normSq ↑(S • z) < 1 := by theorem im_lt_im_S_smul (h : normSq z < 1) : z.im < (S • z).im := by have : z.im < z.im / normSq (z : ℂ) := by have imz : 0 < z.im := im_pos z - apply (lt_div_iff z.normSq_pos).mpr + apply (lt_div_iff₀ z.normSq_pos).mpr nlinarith convert this simp only [ModularGroup.im_smul_eq_div_normSq] @@ -380,7 +380,7 @@ scoped[Modular] notation "𝒟ᵒ" => ModularGroup.fdo open scoped Modular theorem abs_two_mul_re_lt_one_of_mem_fdo (h : z ∈ 𝒟ᵒ) : |2 * z.re| < 1 := by - rw [abs_mul, abs_two, ← lt_div_iff' (zero_lt_two' ℝ)] + rw [abs_mul, abs_two, ← lt_div_iff₀' (zero_lt_two' ℝ)] exact h.2 theorem three_lt_four_mul_im_sq_of_mem_fdo (h : z ∈ 𝒟ᵒ) : 3 < 4 * z.im ^ 2 := by diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index f5d474fad9d31..eb0293caf7695 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -403,7 +403,7 @@ theorem convexBodySum_volume : convert addHaar_smul volume B (convexBodySum K 1) · simp_rw [← Set.preimage_smul_inv₀ (ne_of_gt hB), Set.preimage_setOf_eq, convexBodySumFun, normAtPlace_smul, abs_inv, abs_eq_self.mpr (le_of_lt hB), ← mul_assoc, mul_comm, mul_assoc, - ← Finset.mul_sum, inv_mul_le_iff hB, mul_one] + ← Finset.mul_sum, inv_mul_le_iff₀ hB, mul_one] · rw [abs_pow, ofReal_pow (abs_nonneg _), abs_eq_self.mpr (le_of_lt hB), mixedEmbedding.finrank] · exact this.symm diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index c6726b41347b6..992a53910d440 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -75,7 +75,7 @@ theorem _root_.RingOfIntegers.isPrincipalIdealRing_of_abs_discr_lt ((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 - rw [← Real.sqrt_lt (by positivity) (by positivity), mul_assoc, ← inv_mul_lt_iff' (by positivity), + rw [← Real.sqrt_lt (by positivity) (by positivity), mul_assoc, ← inv_mul_lt_iff₀' (by positivity), mul_inv, ← inv_pow, inv_div, inv_div, mul_assoc, Int.cast_abs] at h rw [← classNumber_eq_one_iff, classNumber, Fintype.card_eq_one_iff] refine ⟨1, fun C ↦ ?_⟩ diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index b93019fd9a8fb..85e0a1770c687 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -173,7 +173,7 @@ theorem abs_discr_ge (h : 1 < finrank ℚ K) : rw [← Algebra.coe_norm_int, ← Int.cast_one, ← Int.cast_abs, Rat.cast_intCast, Int.cast_le] exact Int.one_le_abs (Algebra.norm_ne_zero_iff.mpr h_nz) replace h_bd := le_trans h_nm h_bd - rw [← inv_mul_le_iff (by positivity), inv_div, mul_one, Real.le_sqrt (by positivity) + rw [← inv_mul_le_iff₀ (by positivity), inv_div, mul_one, Real.le_sqrt (by positivity) (by positivity), ← Int.cast_abs, div_pow, mul_pow, ← pow_mul, ← pow_mul] at h_bd refine le_trans ?_ h_bd -- The sequence `a n` is a lower bound for `|discr K|`. We prove below by induction an uniform @@ -277,7 +277,7 @@ theorem rank_le_rankOfDiscrBdd : refine fun h ↦ discr_ne_zero K ?_ rwa [h, Nat.cast_zero, abs_nonpos_iff] at hK have h₂ : 1 < 3 * π / 4 := by - rw [_root_.lt_div_iff (by positivity), ← _root_.div_lt_iff' (by positivity), one_mul] + rw [_root_.lt_div_iff₀ (by positivity), ← _root_.div_lt_iff₀' (by positivity), one_mul] linarith [Real.pi_gt_three] obtain h | h := lt_or_le 1 (finrank ℚ K) · apply le_max_of_le_right diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 2f8a1d3a28cd8..636189cd1f51f 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -304,7 +304,7 @@ theorem exists_unit (w₁ : InfinitePlace K) : _ = 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] + 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 }) diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 2d3a6893edd99..0de565647f9ee 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -401,7 +401,7 @@ private theorem soln_dist_to_a : ‖soln - a‖ = ‖F.eval a‖ / ‖F.derivati tendsto_nhds_unique (newton_seq_dist_tendsto' hnorm) (newton_seq_dist_tendsto hnorm hnsol) private theorem soln_dist_to_a_lt_deriv : ‖soln - a‖ < ‖F.derivative.eval a‖ := by - rw [soln_dist_to_a, div_lt_iff (deriv_norm_pos _), ← sq] <;> assumption + rw [soln_dist_to_a, div_lt_iff₀ (deriv_norm_pos _), ← sq] <;> assumption private theorem soln_unique (z : ℤ_[p]) (hev : F.eval z = 0) (hnlt : ‖z - a‖ < ‖F.derivative.eval a‖) : z = soln := diff --git a/Mathlib/NumberTheory/Rayleigh.lean b/Mathlib/NumberTheory/Rayleigh.lean index d7a2e5f205472..0e29615583c61 100644 --- a/Mathlib/NumberTheory/Rayleigh.lean +++ b/Mathlib/NumberTheory/Rayleigh.lean @@ -63,9 +63,9 @@ private theorem no_collision (hrs : r.IsConjExponent s) : Disjoint {beattySeq r k | k} {beattySeq' s k | k} := by rw [Set.disjoint_left] intro j ⟨k, h₁⟩ ⟨m, h₂⟩ - rw [beattySeq, Int.floor_eq_iff, ← div_le_iff₀ hrs.pos, ← lt_div_iff hrs.pos] at h₁ + rw [beattySeq, Int.floor_eq_iff, ← div_le_iff₀ hrs.pos, ← lt_div_iff₀ hrs.pos] at h₁ rw [beattySeq', sub_eq_iff_eq_add, Int.ceil_eq_iff, Int.cast_add, Int.cast_one, - add_sub_cancel_right, ← div_lt_iff hrs.symm.pos, ← le_div_iff₀ hrs.symm.pos] at h₂ + add_sub_cancel_right, ← div_lt_iff₀ hrs.symm.pos, ← le_div_iff₀ hrs.symm.pos] at h₂ have h₃ := add_lt_add_of_le_of_lt h₁.1 h₂.1 have h₄ := add_lt_add_of_lt_of_le h₁.2 h₂.2 simp_rw [div_eq_inv_mul, ← right_distrib, hrs.inv_add_inv_conj, one_mul] at h₃ h₄ @@ -91,10 +91,10 @@ private theorem hit_or_miss (h : r > 0) : -- for both cases, the candidate is `k = ⌈(j + 1) / r⌉ - 1` cases lt_or_ge ((⌈(j + 1) / r⌉ - 1) * r) j · refine Or.inr ⟨⌈(j + 1) / r⌉ - 1, ?_⟩ - rw [Int.cast_sub, Int.cast_one, lt_div_iff h, sub_add_cancel] + rw [Int.cast_sub, Int.cast_one, lt_div_iff₀ h, sub_add_cancel] exact ⟨‹_›, Int.le_ceil _⟩ · refine Or.inl ⟨⌈(j + 1) / r⌉ - 1, ?_⟩ - rw [beattySeq, Int.floor_eq_iff, Int.cast_sub, Int.cast_one, ← lt_div_iff h, sub_lt_iff_lt_add] + rw [beattySeq, Int.floor_eq_iff, Int.cast_sub, Int.cast_one, ← lt_div_iff₀ h, sub_lt_iff_lt_add] exact ⟨‹_›, Int.ceil_lt_add_one _⟩ /-- Let `0 < r ∈ ℝ` and `j ∈ ℤ`. Then either `j ∈ B'_r` or `B'_r` jumps over `j`. -/ diff --git a/Mathlib/RingTheory/RootsOfUnity/Complex.lean b/Mathlib/RingTheory/RootsOfUnity/Complex.lean index ec56155951282..a11bd30bfad0a 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Complex.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Complex.lean @@ -168,7 +168,7 @@ theorem IsPrimitiveRoot.arg {n : ℕ} {ζ : ℂ} (h : IsPrimitiveRoot ζ n) (hn exact mul_nonpos_of_nonpos_of_nonneg (sub_nonpos.mpr <| mod_cast h.le) (div_nonneg (by simp [Real.pi_pos.le]) <| by simp) rw [← mul_rotate', mul_div_assoc, neg_lt, ← mul_neg, mul_lt_iff_lt_one_right Real.pi_pos, ← - neg_div, ← neg_mul, neg_sub, div_lt_iff, one_mul, sub_mul, sub_lt_comm, ← mul_sub_one] + neg_div, ← neg_mul, neg_sub, div_lt_iff₀, one_mul, sub_mul, sub_lt_comm, ← mul_sub_one] · norm_num exact mod_cast not_le.mp h₂ · exact Nat.cast_pos.mpr hn.bot_lt diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 04cdc12e8032b..31b45d485cdb9 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -106,7 +106,7 @@ theorem unique_topology_of_t2 {t : TopologicalSpace 𝕜} (h₁ : @TopologicalAd -- For that, we use that `𝓑` is balanced : since `‖ξ₀‖ < ε < ‖ξ‖`, we have `‖ξ₀ / ξ‖ ≤ 1`, -- hence `ξ₀ = (ξ₀ / ξ) • ξ ∈ 𝓑` because `ξ ∈ 𝓑`. refine (balancedCore_balanced _).smul_mem ?_ hξ - rw [norm_mul, norm_inv, mul_inv_le_iff (norm_pos_iff.mpr hξ0), mul_one] + rw [norm_mul, norm_inv, mul_inv_le_iff₀ (norm_pos_iff.mpr hξ0), one_mul] exact (hξ₀ε.trans h).le · -- Finally, to show `𝓣₀ ≤ 𝓣`, we simply argue that `id = (fun x ↦ x • 1)` is continuous from -- `(𝕜, 𝓣₀)` to `(𝕜, 𝓣)` because `(•) : (𝕜, 𝓣₀) × (𝕜, 𝓣) → (𝕜, 𝓣)` is continuous. diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index f34d6b2a026ac..fe184b35a5fa5 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -60,9 +60,9 @@ instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by · intro n x h1 h2 rw [hVn] at h1 h2 ⊢ rwa [Circle.coe_mul, Complex.arg_mul x.coe_ne_zero x.coe_ne_zero, - ← two_mul, abs_mul, abs_two, ← lt_div_iff' two_pos, div_div, ← pow_succ] at h2 + ← two_mul, abs_mul, abs_two, ← lt_div_iff₀' two_pos, div_div, ← pow_succ] at h2 apply Set.Ioo_subset_Ioc_self - rw [← two_mul, Set.mem_Ioo, ← abs_lt, abs_mul, abs_two, ← lt_div_iff' two_pos] + rw [← two_mul, Set.mem_Ioo, ← abs_lt, abs_mul, abs_two, ← lt_div_iff₀' two_pos] exact h1.trans_le (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow one_le_two n.succ_ne_zero)) · rw [← Circle.exp_zero, ← isLocalHomeomorph_circleExp.map_nhds_eq 0] diff --git a/Mathlib/Topology/ContinuousMap/Ideals.lean b/Mathlib/Topology/ContinuousMap/Ideals.lean index ead82fefc127c..5edb4350d5274 100644 --- a/Mathlib/Topology/ContinuousMap/Ideals.lean +++ b/Mathlib/Topology/ContinuousMap/Ideals.lean @@ -169,7 +169,7 @@ theorem exists_mul_le_one_eqOn_ge (f : C(X, ℝ≥0)) {c : ℝ≥0} (hc : 0 < c) continuous_toFun := ((map_continuous f).sup <| map_continuous _).inv₀ fun _ => (hc.trans_le le_sup_right).ne' }, fun x => - (inv_mul_le_iff (hc.trans_le le_sup_right)).mpr ((mul_one (f x ⊔ c)).symm ▸ le_sup_left), + (inv_mul_le_iff₀ (hc.trans_le le_sup_right)).mpr ((mul_one (f x ⊔ c)).symm ▸ le_sup_left), fun x hx => by simpa only [coe_const, mul_apply, coe_mk, Pi.inv_apply, Pi.sup_apply, Function.const_apply, sup_eq_left.mpr (Set.mem_setOf.mp hx), ne_eq, Pi.one_apply] From 1a9758b2d278f36febc167f9a830cb9c5f7df799 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 13:49:31 +0000 Subject: [PATCH 150/174] feat: `ENNReal`-valued conjugate exponents (#17353) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `ENNReal.IsConjExponent`, the `ENNReal` analogue of `Real.IsConjExponent` and `NNReal.IsConjExponent`. This will allow stating Hölder's inequality for the L1 and Linfty norms too. From LeanAPAP --- Mathlib/Data/ENNReal/Operations.lean | 5 +- Mathlib/Data/Real/ConjExponents.lean | 125 ++++++++++++++++++++++++++- 2 files changed, 128 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index bf0335bc473f4..473c3d6c65d78 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -256,6 +256,7 @@ section Cancel -- Porting note (#11215): TODO: generalize to `WithTop` /-- An element `a` is `AddLECancellable` if `a + b ≤ a + c` implies `b ≤ c` for all `b` and `c`. This is true in `ℝ≥0∞` for all elements except `∞`. -/ +@[simp] theorem addLECancellable_iff_ne {a : ℝ≥0∞} : AddLECancellable a ↔ a ≠ ∞ := by constructor · rintro h rfl @@ -294,11 +295,13 @@ theorem sub_eq_sInf {a b : ℝ≥0∞} : a - b = sInf { d | a ≤ d + b } := le_antisymm (le_sInf fun _ h => tsub_le_iff_right.mpr h) <| sInf_le <| mem_setOf.2 le_tsub_add /-- This is a special case of `WithTop.coe_sub` in the `ENNReal` namespace -/ -@[simp] theorem coe_sub : (↑(r - p) : ℝ≥0∞) = ↑r - ↑p := WithTop.coe_sub +@[simp, norm_cast] theorem coe_sub : (↑(r - p) : ℝ≥0∞) = ↑r - ↑p := WithTop.coe_sub /-- This is a special case of `WithTop.top_sub_coe` in the `ENNReal` namespace -/ @[simp] theorem top_sub_coe : ∞ - ↑r = ∞ := WithTop.top_sub_coe +@[simp] lemma top_sub (ha : a ≠ ∞) : ∞ - a = ∞ := by lift a to ℝ≥0 using ha; exact top_sub_coe + /-- This is a special case of `WithTop.sub_top` in the `ENNReal` namespace -/ theorem sub_top : a - ∞ = 0 := WithTop.sub_top diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 5323170031fd4..d315cee34de3e 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -18,6 +18,8 @@ analysis, especially when dealing with `L^p` spaces. * `Real.conjExponent`: Conjugate exponent of a real number. * `NNReal.IsConjExponent`: Predicate for two nonnegative real numbers to be conjugate. * `NNReal.conjExponent`: Conjugate exponent of a nonnegative real number. +* `ENNReal.IsConjExponent`: Predicate for two extended nonnegative real numbers to be conjugate. +* `ENNReal.conjExponent`: Conjugate exponent of an extended nonnegative real number. ## TODO @@ -27,7 +29,7 @@ analysis, especially when dealing with `L^p` spaces. noncomputable section -open scoped ENNReal +open scoped ENNReal NNReal namespace Real @@ -116,6 +118,8 @@ lemma one_sub_inv_inv (ha₀ : 0 < a) (ha₁ : a < 1) : (1 - a)⁻¹.IsConjExpon end IsConjExponent +lemma isConjExponent_comm : p.IsConjExponent q ↔ q.IsConjExponent p := ⟨.symm, .symm⟩ + lemma isConjExponent_iff_eq_conjExponent (hp : 1 < p) : p.IsConjExponent q ↔ q = p / (p - 1) := ⟨IsConjExponent.conj_eq, fun h ↦ ⟨hp, by field_simp [h]⟩⟩ @@ -206,6 +210,8 @@ lemma one_sub_inv_inv (ha₀ : a ≠ 0) (ha₁ : a < 1) : (1 - a)⁻¹.IsConjExp end IsConjExponent +lemma isConjExponent_comm : p.IsConjExponent q ↔ q.IsConjExponent p := ⟨.symm, .symm⟩ + lemma isConjExponent_iff_eq_conjExponent (h : 1 < p) : p.IsConjExponent q ↔ q = p / (p - 1) := by rw [← isConjExponent_coe, Real.isConjExponent_iff_eq_conjExponent (mod_cast h), ← coe_inj, NNReal.coe_div, NNReal.coe_sub h.le, coe_one] @@ -220,3 +226,120 @@ protected lemma Real.IsConjExponent.toNNReal {p q : ℝ} (hpq : p.IsConjExponent one_lt := by simpa using hpq.one_lt inv_add_inv_conj := by rw [← toNNReal_inv, ← toNNReal_inv, ← toNNReal_add hpq.inv_nonneg hpq.symm.inv_nonneg, hpq.inv_add_inv_conj, toNNReal_one] + +namespace ENNReal + +/-- Two extended nonnegative real exponents `p, q` are conjugate and satisfy the equality +`1/p + 1/q = 1`. This condition shows up in many theorems in analysis, notably related to `L^p` +norms. Note that we permit one of the exponents to be `∞` and the other `1`. -/ +@[mk_iff] +structure IsConjExponent (p q : ℝ≥0∞) : Prop where + inv_add_inv_conj : p⁻¹ + q⁻¹ = 1 + +/-- The conjugate exponent of `p` is `q = 1 + (p - 1)⁻¹`, so that `1/p + 1/q = 1`. -/ +noncomputable def conjExponent (p : ℝ≥0∞) : ℝ≥0∞ := 1 + (p - 1)⁻¹ + +lemma coe_conjExponent {p : ℝ≥0} (hp : 1 < p) : p.conjExponent = conjExponent p := by + rw [NNReal.conjExponent, conjExponent] + norm_cast + rw [← coe_inv (tsub_pos_of_lt hp).ne'] + norm_cast + field_simp [(tsub_pos_of_lt hp).ne'] + rw [tsub_add_cancel_of_le hp.le] + +variable {a b p q : ℝ≥0∞} (h : p.IsConjExponent q) + +@[simp, norm_cast] lemma isConjExponent_coe {p q : ℝ≥0} : + IsConjExponent p q ↔ p.IsConjExponent q := by + simp only [isConjExponent_iff, NNReal.isConjExponent_iff] + refine ⟨fun h ↦ ⟨?_, ?_⟩, ?_⟩ + · simpa using (ENNReal.lt_add_right (fun hp ↦ by simp [hp] at h) <| by simp).trans_eq h + · rw [← coe_inv, ← coe_inv] at h + norm_cast at h + all_goals rintro rfl; simp at h + · rintro ⟨hp, h⟩ + rw [← coe_inv (zero_lt_one.trans hp).ne', ← coe_inv, ← coe_add, h, coe_one] + rintro rfl + simp [hp.ne'] at h + +alias ⟨_, _root_.NNReal.IsConjExponent.coe_ennreal⟩ := isConjExponent_coe + +namespace IsConjExponent + +protected lemma conjExponent (hp : 1 ≤ p) : p.IsConjExponent (conjExponent p) := by + have : p ≠ 0 := (zero_lt_one.trans_le hp).ne' + rw [isConjExponent_iff, conjExponent, add_comm] + 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 + obtain rfl | hp := eq_or_ne p ∞ + · simp + calc + 1 + (p - 1)⁻¹ = (p - 1 + 1) / (p - 1) := by + rw [ENNReal.add_div, ENNReal.div_self ((tsub_pos_of_lt hp₁).ne') (sub_ne_top hp), one_div] + _ = (1 - p⁻¹)⁻¹ := by + rw [tsub_add_cancel_of_le, ← inv_eq_iff_eq_inv, div_eq_mul_inv, ENNReal.mul_inv, inv_inv, + ENNReal.mul_sub, ENNReal.inv_mul_cancel, mul_one] <;> simp [*] + +section +include h + +@[symm] +protected lemma symm : q.IsConjExponent p where + inv_add_inv_conj := by simpa [add_comm] using h.inv_add_inv_conj + +lemma one_le : 1 ≤ p := ENNReal.inv_le_one.1 <| by + rw [← add_zero p⁻¹, ← h.inv_add_inv_conj]; gcongr; positivity + +lemma pos : 0 < p := zero_lt_one.trans_le h.one_le +lemma ne_zero : p ≠ 0 := h.pos.ne' + +lemma one_sub_inv : 1 - p⁻¹ = q⁻¹ := + ENNReal.sub_eq_of_eq_add_rev' one_ne_top h.inv_add_inv_conj.symm + +lemma conjExponent_eq : conjExponent p = q := by + have hp : 1 ≤ p := h.one_le + have : p⁻¹ ≠ ∞ := by simpa using h.ne_zero + simpa [ENNReal.add_right_inj, *] using + (IsConjExponent.conjExponent hp).inv_add_inv_conj.trans h.inv_add_inv_conj.symm + +lemma conj_eq : q = 1 + (p - 1)⁻¹ := h.conjExponent_eq.symm + +lemma mul_eq_add : p * q = p + q := by + obtain rfl | hp := eq_or_ne p ∞ + · simp [h.symm.ne_zero] + obtain rfl | hq := eq_or_ne q ∞ + · simp [h.ne_zero] + rw [← mul_one (_ * _), ← h.inv_add_inv_conj, mul_add, mul_right_comm, + ENNReal.mul_inv_cancel h.ne_zero hp, one_mul, mul_assoc, + ENNReal.mul_inv_cancel h.symm.ne_zero hq, mul_one, add_comm] + +lemma div_conj_eq_sub_one : p / q = p - 1 := by + obtain rfl | hq := eq_or_ne q ∞ + · simp [h.symm.conj_eq] + 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] + +end + +protected lemma inv_inv (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ where + inv_add_inv_conj := by simpa only [inv_inv] using hab + +lemma inv_one_sub_inv (ha : a ≤ 1) : a⁻¹.IsConjExponent (1 - a)⁻¹ := + .inv_inv <| add_tsub_cancel_of_le ha + +lemma one_sub_inv_inv (ha : a ≤ 1) : (1 - a)⁻¹.IsConjExponent a⁻¹ := (inv_one_sub_inv ha).symm + +lemma top_one : IsConjExponent ∞ 1 := ⟨by simp⟩ +lemma one_top : IsConjExponent 1 ∞ := ⟨by simp⟩ + +end IsConjExponent + +lemma isConjExponent_comm : p.IsConjExponent q ↔ q.IsConjExponent p := ⟨.symm, .symm⟩ + +lemma isConjExponent_iff_eq_conjExponent (hp : 1 ≤ p) : p.IsConjExponent q ↔ q = 1 + (p - 1)⁻¹ := + ⟨fun h ↦ h.conj_eq, by rintro rfl; exact .conjExponent hp⟩ + +end ENNReal From d369bba55fcefeb0d7600e9c2e4856ac9bbece72 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 3 Oct 2024 13:49:32 +0000 Subject: [PATCH 151/174] chore: update Mathlib dependencies 2024-10-03 (#17373) 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 efc4091d6b8c8..027d26b4ddaa2 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "34e690ec07f6f6375668adba5a16d0d723226c2c", + "rev": "f274aed7ae8d1addd3e70adaf3183ccc6e1ed43d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 915e456f649faa69c395e65f18d3db713717112d Mon Sep 17 00:00:00 2001 From: D-Thomine <100795491+D-Thomine@users.noreply.github.com> Date: Thu, 3 Oct 2024 14:46:32 +0000 Subject: [PATCH 152/174] feat(UniformSpace.Basic): add ball_preimage (#17375) Simple lemma about the preimage of a ball and the ball of the preimage of an entourage. --- Mathlib/Topology/UniformSpace/Basic.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 3e05609edfe68..4d5a67fa923b1 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -892,7 +892,6 @@ lemma DenseRange.iUnion_uniformity_ball {ι : Type*} {xs : ι → α} ### Uniformity bases -/ - /-- Open elements of `𝓤 α` form a basis of `𝓤 α`. -/ theorem uniformity_hasBasis_open : HasBasis (𝓤 α) (fun V : Set (α × α) => V ∈ 𝓤 α ∧ IsOpen V) id := hasBasis_self.2 fun s hs => @@ -1117,6 +1116,11 @@ theorem uniformity_comap {_ : UniformSpace β} (f : α → β) : 𝓤[UniformSpace.comap f ‹_›] = comap (Prod.map f f) (𝓤 β) := rfl +lemma ball_preimage {f : α → β} {U : Set (β × β)} {x : α} : + UniformSpace.ball x (Prod.map f f ⁻¹' U) = f ⁻¹' UniformSpace.ball (f x) U := by + ext : 1 + simp only [UniformSpace.ball, mem_preimage, Prod.map_apply] + @[simp] theorem uniformSpace_comap_id {α : Type*} : UniformSpace.comap (id : α → α) = id := by ext : 2 From 4f0309fc7dfd2e07731434f6bb1dd6ef2e4cd85b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 3 Oct 2024 15:18:46 +0000 Subject: [PATCH 153/174] style(GroupTheory,RingTheory/Congruence/Basic): use `where`, fix whitespace (#17362) --- Mathlib/GroupTheory/Congruence/Basic.lean | 1 - Mathlib/RingTheory/Congruence/Basic.lean | 18 ++++++++---------- 2 files changed, 8 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index a3011ab4ba7d6..59ad69ed64e0d 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -110,7 +110,6 @@ variable [Mul M] [Mul N] [Mul P] (c : Con M) instance : Inhabited (Con M) := ⟨conGen EmptyRelation⟩ --- Porting note: upgraded to FunLike /-- A coercion from a congruence relation to its underlying binary relation. -/ @[to_additive "A coercion from an additive congruence relation to its underlying binary relation."] instance : FunLike (Con M) M (M → Prop) where diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index c6c46c89d1b9b..6d8463a131488 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -67,16 +67,15 @@ section Basic variable [Add R] [Mul R] (c : RingCon R) --- Porting note: upgrade to `FunLike` /-- A coercion from a congruence relation to its underlying binary relation. -/ -instance : FunLike (RingCon R) R (R → Prop) := - { coe := fun c => c.r, - coe_injective' := fun x y h => by - rcases x with ⟨⟨x, _⟩, _⟩ - rcases y with ⟨⟨y, _⟩, _⟩ - congr! - rw [Setoid.ext_iff,(show x.Rel = y.Rel from h)] - simp} +instance : FunLike (RingCon R) R (R → Prop) where + coe := fun c => c.r + coe_injective' x y h := by + rcases x with ⟨⟨x, _⟩, _⟩ + rcases y with ⟨⟨y, _⟩, _⟩ + congr! + rw [Setoid.ext_iff, (show x.Rel = y.Rel from h)] + simp theorem rel_eq_coe : c.r = c := rfl @@ -419,7 +418,6 @@ end Quotient The API in this section is copied from `Mathlib/GroupTheory/Congruence.lean` -/ - section Lattice variable [Add R] [Mul R] From fa42460aae66e4c5f6cf6a59db86db0383406022 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 3 Oct 2024 15:18:47 +0000 Subject: [PATCH 154/174] feat: golf using `module`/`match_scalars` throughout the library (#17365) 50 sample uses of the new (#16593) `module` and `match_scalars` tactics. --- Mathlib/Algebra/Polynomial/Module/Basic.lean | 10 ++-- Mathlib/Analysis/Analytic/Meromorphic.lean | 8 +-- .../Analysis/Calculus/FDeriv/Symmetric.lean | 57 +++++++------------ .../Analysis/Calculus/UniformLimitsDeriv.lean | 4 +- Mathlib/Analysis/Convex/Basic.lean | 14 ++--- Mathlib/Analysis/Convex/Combination.lean | 7 ++- Mathlib/Analysis/Convex/Jensen.lean | 7 +-- Mathlib/Analysis/Convex/Join.lean | 22 +++---- Mathlib/Analysis/Convex/Side.lean | 23 +++----- Mathlib/Analysis/Convex/Star.lean | 5 +- Mathlib/Analysis/Convex/StoneSeparation.lean | 16 ++---- Mathlib/Analysis/Convex/Strict.lean | 13 ++--- Mathlib/Analysis/NormedSpace/Connected.lean | 10 ++-- .../Euclidean/Angle/Oriented/Rotation.lean | 19 +++---- Mathlib/LinearAlgebra/LinearIndependent.lean | 26 ++++----- .../LinearAlgebra/QuadraticForm/Basic.lean | 17 +++--- Mathlib/LinearAlgebra/Ray.lean | 12 ++-- .../NumberTheory/LSeries/AbstractFuncEq.lean | 11 ++-- 18 files changed, 118 insertions(+), 163 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index c5e69f5e9a19f..f21958832dbda 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -276,8 +276,7 @@ theorem eval_smul (p : R[X]) (q : PolynomialModule R M) (r : R) : intro i m induction p using Polynomial.induction_on' with | h_add _ _ e₁ e₂ => rw [add_smul, map_add, Polynomial.eval_add, e₁, e₂, add_smul] - | h_monomial => rw [monomial_smul_single, eval_single, Polynomial.eval_monomial, eval_single, - smul_comm, ← smul_smul, pow_add, mul_smul] + | h_monomial => simp only [monomial_smul_single, Polynomial.eval_monomial, eval_single]; module @[simp] theorem eval_map (f : M →ₗ[R] M') (q : PolynomialModule R M) (r : R) : @@ -287,7 +286,8 @@ theorem eval_map (f : M →ₗ[R] M') (q : PolynomialModule R M) (r : R) : · intro f g e₁ e₂ simp_rw [map_add, e₁, e₂] · intro i m - rw [map_single, eval_single, eval_single, f.map_smul, ← map_pow, algebraMap_smul] + simp only [map_single, eval_single, f.map_smul] + module @[simp] theorem eval_map' (f : M →ₗ[R] M) (q : PolynomialModule R M) (r : R) : @@ -324,8 +324,8 @@ theorem comp_eval (p : R[X]) (q : PolynomialModule R M) (r : R) : · intro _ _ e₁ e₂ simp_rw [map_add, e₁, e₂] · intro i m - rw [LinearMap.comp_apply, comp_single, eval_single, eval_smul, eval_single, pow_zero, one_smul, - Polynomial.eval_pow] + rw [LinearMap.comp_apply, comp_single, eval_single, eval_smul, eval_single, eval_pow] + module theorem comp_smul (p p' : R[X]) (q : PolynomialModule R M) : comp p (p' • q) = p'.comp p • comp p q := by diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index 4c3349d87d078..be73e28bbbfe1 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -59,8 +59,8 @@ lemma smul {f : 𝕜 → 𝕜} {g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f rcases hg with ⟨n, hg⟩ refine ⟨m + n, ?_⟩ convert hf.smul hg using 2 with z - rw [smul_eq_mul, ← mul_smul, mul_assoc, mul_comm (f z), ← mul_assoc, pow_add, - ← smul_eq_mul (a' := f z), smul_assoc, Pi.smul_apply'] + rw [Pi.smul_apply', smul_eq_mul] + module lemma mul {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f * g) x := @@ -227,8 +227,8 @@ lemma iff_eventuallyEq_zpow_smul_analyticAt {f : 𝕜 → E} {x : 𝕜} : Meromo ∃ (n : ℤ) (g : 𝕜 → E), AnalyticAt 𝕜 g x ∧ ∀ᶠ z in 𝓝[≠] x, f z = (z - x) ^ n • g z := by refine ⟨fun ⟨n, hn⟩ ↦ ⟨-n, _, ⟨hn, eventually_nhdsWithin_iff.mpr ?_⟩⟩, ?_⟩ · filter_upwards with z hz - rw [← mul_smul, ← zpow_natCast, ← zpow_add₀ (sub_ne_zero.mpr hz), neg_add_cancel, - zpow_zero, one_smul] + match_scalars + field_simp [sub_ne_zero.mpr hz] · refine fun ⟨n, g, hg_an, hg_eq⟩ ↦ MeromorphicAt.congr ?_ (EventuallyEq.symm hg_eq) exact (((MeromorphicAt.id x).sub (.const _ x)).zpow _).smul hg_an.meromorphicAt diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index 1835d8c1d6fa5..399dada358cca 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -97,8 +97,8 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) rw [← smul_smul] apply s_conv.interior.add_smul_mem this _ ht rw [add_assoc] at hw - rw [add_assoc, ← smul_add] - exact s_conv.add_smul_mem_interior xs hw ⟨hpos, h_lt_1.le⟩ + convert s_conv.add_smul_mem_interior xs hw ⟨hpos, h_lt_1.le⟩ using 1 + module -- define a function `g` on `[0,1]` (identified with `[v, v + w]`) such that `g 1 - g 0` is the -- quantity to be estimated. We will check that its derivative is given by an explicit -- expression `g'`, that we can bound. Then the desired bound for `g 1 - g 0` follows from the @@ -139,14 +139,14 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) calc ‖g' t‖ = ‖(f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)) (h • w)‖ := by rw [hg'] - have : h * (t * h) = t * (h * h) := by ring - simp only [ContinuousLinearMap.coe_sub', ContinuousLinearMap.map_add, pow_two, - ContinuousLinearMap.add_apply, Pi.smul_apply, smul_sub, smul_add, smul_smul, ← sub_sub, - ContinuousLinearMap.coe_smul', Pi.sub_apply, ContinuousLinearMap.map_smul, this] + congrm ‖?_‖ + simp only [ContinuousLinearMap.sub_apply, ContinuousLinearMap.add_apply, + ContinuousLinearMap.smul_apply, map_add, map_smul] + module _ ≤ ‖f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)‖ * ‖h • w‖ := (ContinuousLinearMap.le_opNorm _ _) _ ≤ ε * ‖h • v + (t * h) • w‖ * ‖h • w‖ := by - apply mul_le_mul_of_nonneg_right _ (norm_nonneg _) + gcongr have H : x + h • v + (t * h) • w ∈ Metric.ball x δ ∩ interior s := by refine ⟨?_, xt_mem t ⟨ht.1, ht.2.le⟩⟩ rw [add_assoc, add_mem_ball_iff_norm] @@ -157,7 +157,7 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) apply (norm_add_le _ _).trans gcongr simp only [norm_smul, Real.norm_eq_abs, abs_mul, abs_of_nonneg, ht.1, hpos.le, mul_assoc] - exact mul_le_of_le_one_left (mul_nonneg hpos.le (norm_nonneg _)) ht.2.le + exact mul_le_of_le_one_left (by positivity) ht.2.le _ = ε * ((‖v‖ + ‖w‖) * ‖w‖) * h ^ 2 := by simp only [norm_smul, Real.norm_eq_abs, abs_mul, abs_of_nonneg, hpos.le]; ring -- conclude using the mean value inequality @@ -169,8 +169,8 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) simp only [g, Nat.one_ne_zero, add_zero, one_mul, zero_div, zero_mul, sub_zero, zero_smul, Ne, not_false_iff, zero_pow, reduceCtorEq] abel - · simp only [Real.norm_eq_abs, abs_mul, add_nonneg (norm_nonneg v) (norm_nonneg w), abs_of_nonneg, - hpos.le, mul_assoc, norm_nonneg, abs_pow] + · simp (discharger := positivity) only [Real.norm_eq_abs, abs_mul, abs_of_nonneg, abs_pow] + ring /-- One can get `f'' v w` as the limit of `h ^ (-2)` times the alternate sum of the values of `f` along the vertices of a quadrilateral with sides `h v` and `h w` based at `x`. @@ -183,40 +183,27 @@ theorem Convex.isLittleO_alternate_sum_square {v w : E} (h4v : x + (4 : ℝ) • fun h => h ^ 2 := by have A : (1 : ℝ) / 2 ∈ Ioc (0 : ℝ) 1 := ⟨by norm_num, by norm_num⟩ have B : (1 : ℝ) / 2 ∈ Icc (0 : ℝ) 1 := ⟨by norm_num, by norm_num⟩ - have C : ∀ w : E, (2 : ℝ) • w = 2 • w := fun w => by simp only [two_smul] have h2v2w : x + (2 : ℝ) • v + (2 : ℝ) • w ∈ interior s := by convert s_conv.interior.add_smul_sub_mem h4v h4w B using 1 - simp only [smul_sub, smul_smul, one_div, add_sub_add_left_eq_sub, mul_add, add_smul] - norm_num - simp only [show (4 : ℝ) = (2 : ℝ) + (2 : ℝ) by norm_num, _root_.add_smul] - abel + module have h2vww : x + (2 • v + w) + w ∈ interior s := by convert h2v2w using 1 - simp only [two_smul] - abel + module have h2v : x + (2 : ℝ) • v ∈ interior s := by convert s_conv.add_smul_sub_mem_interior xs h4v A using 1 - simp only [smul_smul, one_div, add_sub_cancel_left, add_right_inj] - norm_num + module have h2w : x + (2 : ℝ) • w ∈ interior s := by convert s_conv.add_smul_sub_mem_interior xs h4w A using 1 - simp only [smul_smul, one_div, add_sub_cancel_left, add_right_inj] - norm_num + module have hvw : x + (v + w) ∈ interior s := by convert s_conv.add_smul_sub_mem_interior xs h2v2w A using 1 - simp only [smul_smul, one_div, add_sub_cancel_left, add_right_inj, smul_add, smul_sub] - norm_num - abel + module have h2vw : x + (2 • v + w) ∈ interior s := by convert s_conv.interior.add_smul_sub_mem h2v h2v2w B using 1 - simp only [smul_add, smul_sub, smul_smul, ← C] - norm_num - abel + module have hvww : x + (v + w) + w ∈ interior s := by convert s_conv.interior.add_smul_sub_mem h2w h2v2w B using 1 - rw [one_div, add_sub_add_right_eq_sub, add_sub_cancel_left, inv_smul_smul₀ two_ne_zero, - two_smul] - abel + module have TA1 := s_conv.taylor_approx_two_segment hf xs hx h2vw h2vww have TA2 := s_conv.taylor_approx_two_segment hf xs hx hvw hvww convert TA1.sub TA2 using 1 @@ -245,11 +232,9 @@ theorem Convex.second_derivative_within_at_symmetric_of_mem_interior {v w : E} apply C.congr' _ _ · filter_upwards [self_mem_nhdsWithin] intro h (hpos : 0 < h) - rw [← one_smul ℝ (f'' w v - f'' v w), smul_smul, smul_smul] - congr 1 - field_simp [LT.lt.ne' hpos] + match_scalars <;> field_simp · filter_upwards [self_mem_nhdsWithin] with h (hpos : 0 < h) - field_simp [LT.lt.ne' hpos, SMul.smul] + field_simp simpa only [sub_eq_zero] using isLittleO_const_const_iff.1 B end @@ -298,8 +283,8 @@ theorem Convex.second_derivative_within_at_symmetric {s : Set E} (s_conv : Conve s_conv.second_derivative_within_at_symmetric_of_mem_interior hf xs hx (ts w) (ts v) simp only [ContinuousLinearMap.map_add, ContinuousLinearMap.map_smul, smul_add, smul_smul, ContinuousLinearMap.add_apply, Pi.smul_apply, ContinuousLinearMap.coe_smul', C] at this - rw [add_assoc, add_assoc, add_right_inj, add_left_comm, add_right_inj, add_right_inj, mul_comm] - at this + have : (t v * t w) • (f'' v) w = (t v * t w) • (f'' w) v := by + linear_combination (norm := module) this apply smul_right_injective F _ this simp [(tpos v).ne', (tpos w).ne'] diff --git a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean index df41e2606b7e8..434945a20c58c 100644 --- a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean +++ b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean @@ -356,8 +356,8 @@ theorem hasFDerivAt_of_tendstoUniformlyOnFilter [NeBot l] apply ((this ε hε).filter_mono curry_le_prod).mono intro n hn rw [dist_eq_norm] at hn ⊢ - rw [← smul_sub] at hn - rwa [sub_zero] + convert hn using 2 + module · -- (Almost) the definition of the derivatives rw [Metric.tendsto_nhds] intro ε hε diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index f07e88d57f245..1eb8cdb030baa 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -8,6 +8,8 @@ import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Analysis.Convex.Star import Mathlib.LinearAlgebra.AffineSpace.AffineSubspace +import Mathlib.Tactic.FieldSimp +import Mathlib.Tactic.NoncommRing /-! # Convex sets and functions in vector spaces @@ -158,8 +160,7 @@ theorem convex_segment (x y : E) : Convex 𝕜 [x -[𝕜] y] := by ⟨a * ap + b * aq, a * bp + b * bq, add_nonneg (mul_nonneg ha hap) (mul_nonneg hb haq), add_nonneg (mul_nonneg ha hbp) (mul_nonneg hb hbq), ?_, ?_⟩ · rw [add_add_add_comm, ← mul_add, ← mul_add, habp, habq, mul_one, mul_one, hab] - · simp_rw [add_smul, mul_smul, smul_add] - exact add_add_add_comm _ _ _ _ + · match_scalars <;> noncomm_ring theorem Convex.linear_image (hs : Convex 𝕜 s) (f : E →ₗ[𝕜] F) : Convex 𝕜 (f '' s) := by rintro _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩ a b ha hb hab @@ -406,8 +407,8 @@ theorem convex_openSegment (a b : E) : Convex 𝕜 (openSegment 𝕜 a b) := by rw [convex_iff_openSegment_subset] rintro p ⟨ap, bp, hap, hbp, habp, rfl⟩ q ⟨aq, bq, haq, hbq, habq, rfl⟩ z ⟨a, b, ha, hb, hab, rfl⟩ refine ⟨a * ap + b * aq, a * bp + b * bq, by positivity, by positivity, ?_, ?_⟩ - · rw [add_add_add_comm, ← mul_add, ← mul_add, habp, habq, mul_one, mul_one, hab] - · simp_rw [add_smul, mul_smul, smul_add, add_add_add_comm] + · linear_combination (norm := noncomm_ring) a * habp + b * habq + hab + · module end StrictOrderedCommSemiring @@ -425,8 +426,7 @@ theorem convex_vadd (a : E) : Convex 𝕜 (a +ᵥ s) ↔ Convex 𝕜 s := theorem Convex.add_smul_mem (hs : Convex 𝕜 s) {x y : E} (hx : x ∈ s) (hy : x + y ∈ s) {t : 𝕜} (ht : t ∈ Icc (0 : 𝕜) 1) : x + t • y ∈ s := by - have h : x + t • y = (1 - t) • x + t • (x + y) := by - rw [smul_add, ← add_assoc, ← add_smul, sub_add_cancel, one_smul] + have h : x + t • y = (1 - t) • x + t • (x + y) := by match_scalars <;> noncomm_ring rw [h] exact hs hx hy (sub_nonneg_of_le ht.2) ht.1 (sub_add_cancel _ _) @@ -515,7 +515,7 @@ theorem Convex.exists_mem_add_smul_eq (h : Convex 𝕜 s) {x y : E} {p q : 𝕜} · replace hpq : 0 < p + q := (add_nonneg hp hq).lt_of_ne' (mt (add_eq_zero_iff_of_nonneg hp hq).1 hpq) refine ⟨_, convex_iff_div.1 h hx hy hp hq hpq, ?_⟩ - simp only [smul_add, smul_smul, mul_div_cancel₀ _ hpq.ne'] + match_scalars <;> field_simp theorem Convex.add_smul (h_conv : Convex 𝕜 s) {p q : 𝕜} (hp : 0 ≤ p) (hq : 0 ≤ q) : (p + q) • s = p • s + q • s := (add_smul_subset _ _ _).antisymm <| by diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 1fc98bcd1e819..399464ada7406 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -50,7 +50,8 @@ theorem Finset.centerMass_empty : (∅ : Finset ι).centerMass w z = 0 := by theorem Finset.centerMass_pair (hne : i ≠ j) : ({i, j} : Finset ι).centerMass w z = (w i / (w i + w j)) • z i + (w j / (w i + w j)) • z j := by - simp only [centerMass, sum_pair hne, smul_add, (mul_smul _ _ _).symm, div_eq_inv_mul] + simp only [centerMass, sum_pair hne] + module variable {w} @@ -63,7 +64,9 @@ theorem Finset.centerMass_insert (ha : i ∉ t) (hw : ∑ j ∈ t, w j ≠ 0) : rw [div_mul_eq_mul_div, mul_inv_cancel₀ hw, one_div] theorem Finset.centerMass_singleton (hw : w i ≠ 0) : ({i} : Finset ι).centerMass w z = z i := by - rw [centerMass, sum_singleton, sum_singleton, ← mul_smul, inv_mul_cancel₀ hw, one_smul] + rw [centerMass, sum_singleton, sum_singleton] + match_scalars + field_simp @[simp] lemma Finset.centerMass_neg_left : t.centerMass (-w) z = t.centerMass w z := by simp [centerMass, inv_neg] diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 20e5b017a492d..93f195e014393 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -118,11 +118,10 @@ lemma StrictConvexOn.map_sum_lt (hf : StrictConvexOn 𝕜 s f) (h₀ : ∀ i ∈ have := h₀ k <| by simp let c := w j + w k have hc : w j / c + w k / c = 1 := by field_simp - have hcj : c * (w j / c) = w j := by field_simp - have hck : c * (w k / c) = w k := by field_simp calc f (w j • p j + (w k • p k + ∑ x ∈ u, w x • p x)) _ = f (c • ((w j / c) • p j + (w k / c) • p k) + ∑ x ∈ u, w x • p x) := by - rw [smul_add, ← mul_smul, ← mul_smul, hcj, hck, add_assoc] + congrm f ?_ + match_scalars <;> field_simp _ ≤ c • f ((w j / c) • p j + (w k / c) • p k) + ∑ x ∈ u, w x • f (p x) := -- apply the usual Jensen's inequality wrt the weighted average of the two distinguished -- points and all the other points @@ -134,7 +133,7 @@ lemma StrictConvexOn.map_sum_lt (hf : StrictConvexOn 𝕜 s f) (h₀ : ∀ i ∈ -- then apply the definition of strict convexity for the two distinguished points gcongr; refine hf.2 (hmem _ <| by simp) (hmem _ <| by simp) hjk ?_ ?_ hc <;> positivity _ = (w j • f (p j) + w k • f (p k)) + ∑ x ∈ u, w x • f (p x) := by - rw [smul_add, ← mul_smul, ← mul_smul, hcj, hck] + match_scalars <;> field_simp _ = w j • f (p j) + (w k • f (p k) + ∑ x ∈ u, w x • f (p x)) := by abel_nf /-- Concave **strict Jensen inequality**. diff --git a/Mathlib/Analysis/Convex/Join.lean b/Mathlib/Analysis/Convex/Join.lean index a8641487cfccd..590ebaa4aa614 100644 --- a/Mathlib/Analysis/Convex/Join.lean +++ b/Mathlib/Analysis/Convex/Join.lean @@ -113,19 +113,13 @@ theorem convexJoin_assoc_aux (s t u : Set E) : rintro _ ⟨z, ⟨x, hx, y, hy, a₁, b₁, ha₁, hb₁, hab₁, rfl⟩, z, hz, a₂, b₂, ha₂, hb₂, hab₂, rfl⟩ obtain rfl | hb₂ := hb₂.eq_or_lt · refine ⟨x, hx, y, ⟨y, hy, z, hz, left_mem_segment 𝕜 _ _⟩, a₁, b₁, ha₁, hb₁, hab₁, ?_⟩ - rw [add_zero] at hab₂ - rw [hab₂, one_smul, zero_smul, add_zero] - have ha₂b₁ : 0 ≤ a₂ * b₁ := mul_nonneg ha₂ hb₁ - have hab : 0 < a₂ * b₁ + b₂ := add_pos_of_nonneg_of_pos ha₂b₁ hb₂ + linear_combination (norm := module) congr(-$hab₂ • (a₁ • x + b₁ • y)) refine ⟨x, hx, (a₂ * b₁ / (a₂ * b₁ + b₂)) • y + (b₂ / (a₂ * b₁ + b₂)) • z, - ⟨y, hy, z, hz, _, _, ?_, ?_, ?_, rfl⟩, - a₂ * a₁, a₂ * b₁ + b₂, mul_nonneg ha₂ ha₁, hab.le, ?_, ?_⟩ - · exact div_nonneg ha₂b₁ hab.le - · exact div_nonneg hb₂.le hab.le - · rw [← add_div, div_self hab.ne'] - · rw [← add_assoc, ← mul_add, hab₁, mul_one, hab₂] - · simp_rw [smul_add, ← mul_smul, mul_div_cancel₀ _ hab.ne', add_assoc] + ⟨y, hy, z, hz, _, _, by positivity, by positivity, by field_simp, rfl⟩, + a₂ * a₁, a₂ * b₁ + b₂, by positivity, by positivity, ?_, ?_⟩ + · linear_combination a₂ * hab₁ + hab₂ + · match_scalars <;> field_simp theorem convexJoin_assoc (s t u : Set E) : convexJoin 𝕜 (convexJoin 𝕜 s t) u = convexJoin 𝕜 s (convexJoin 𝕜 t u) := by @@ -155,9 +149,9 @@ protected theorem Convex.convexJoin (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) : rcases hs.exists_mem_add_smul_eq hx₁ hx₂ (mul_nonneg hp ha₁) (mul_nonneg hq ha₂) with ⟨x, hxs, hx⟩ rcases ht.exists_mem_add_smul_eq hy₁ hy₂ (mul_nonneg hp hb₁) (mul_nonneg hq hb₂) with ⟨y, hyt, hy⟩ refine ⟨_, hxs, _, hyt, p * a₁ + q * a₂, p * b₁ + q * b₂, ?_, ?_, ?_, ?_⟩ <;> try positivity - · rwa [add_add_add_comm, ← mul_add, ← mul_add, hab₁, hab₂, mul_one, mul_one] - · rw [hx, hy, add_add_add_comm] - simp only [smul_add, smul_smul] + · linear_combination p * hab₁ + q * hab₂ + hpq + · rw [hx, hy] + module protected theorem Convex.convexHull_union (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hs₀ : s.Nonempty) (ht₀ : t.Nonempty) : convexHull 𝕜 (s ∪ t) = convexJoin 𝕜 s t := diff --git a/Mathlib/Analysis/Convex/Side.lean b/Mathlib/Analysis/Convex/Side.lean index 4bedc05dc363e..f4c36915e34ea 100644 --- a/Mathlib/Analysis/Convex/Side.lean +++ b/Mathlib/Analysis/Convex/Side.lean @@ -333,12 +333,8 @@ theorem _root_.Wbtw.wOppSide₁₃ {s : AffineSubspace R P} {x y z : P} (h : Wbt rcases ht0.lt_or_eq with (ht0' | rfl); swap · rw [lineMap_apply_zero]; simp refine Or.inr (Or.inr ⟨1 - t, t, sub_pos.2 ht1', ht0', ?_⟩) - -- TODO: after lean4#2336 "simp made no progress feature" - -- had to add `_` to several lemmas here. Not sure why! - simp_rw [lineMap_apply _, vadd_vsub_assoc _, vsub_vadd_eq_vsub_sub _, - ← neg_vsub_eq_vsub_rev z x, vsub_self _, zero_sub, ← neg_one_smul R (z -ᵥ x), - ← add_smul, smul_neg, ← neg_smul, smul_smul] - ring_nf + rw [lineMap_apply, vadd_vsub_assoc, vsub_vadd_eq_vsub_sub, ← neg_vsub_eq_vsub_rev z, vsub_self] + module theorem _root_.Wbtw.wOppSide₃₁ {s : AffineSubspace R P} {x y z : P} (h : Wbtw R x y z) (hy : y ∈ s) : s.WOppSide z x := @@ -411,9 +407,9 @@ theorem wOppSide_iff_exists_left {s : AffineSubspace R P} {x y p₁ : P} (h : p exact SameRay.zero_right _ · refine Or.inr ⟨(-r₁ / r₂) • (p₁ -ᵥ p₁') +ᵥ p₂', s.smul_vsub_vadd_mem _ h hp₁' hp₂', Or.inr (Or.inr ⟨r₁, r₂, hr₁, hr₂, ?_⟩)⟩ - rw [vadd_vsub_assoc, smul_add, ← hr, smul_smul, neg_div, mul_neg, - mul_div_cancel₀ _ hr₂.ne.symm, neg_smul, neg_add_eq_sub, ← smul_sub, - vsub_sub_vsub_cancel_right] + rw [vadd_vsub_assoc, ← vsub_sub_vsub_cancel_right x p₁ p₁'] + linear_combination (norm := match_scalars <;> field_simp) hr + ring · rintro (h' | ⟨h₁, h₂, h₃⟩) · exact wOppSide_of_left_mem y h' · exact ⟨p₁, h, h₁, h₂, h₃⟩ @@ -584,15 +580,14 @@ theorem wOppSide_iff_exists_wbtw {s : AffineSubspace R P} {x y : P} : · refine ⟨lineMap x y (r₂ / (r₁ + r₂)), ?_, ?_⟩ · have : (r₂ / (r₁ + r₂)) • (y -ᵥ p₂ + (p₂ -ᵥ p₁) - (x -ᵥ p₁)) + (x -ᵥ p₁) = (r₂ / (r₁ + r₂)) • (p₂ -ᵥ p₁) := by - rw [add_comm (y -ᵥ p₂), smul_sub, smul_add, add_sub_assoc, add_assoc, add_right_eq_self, - div_eq_inv_mul, ← neg_vsub_eq_vsub_rev, smul_neg, ← smul_smul, ← h, smul_smul, ← neg_smul, - ← sub_smul, ← div_eq_inv_mul, ← div_eq_inv_mul, ← neg_div, ← sub_div, sub_eq_add_neg, - ← neg_add, neg_div, div_self (Left.add_pos hr₁ hr₂).ne.symm, neg_one_smul, neg_add_cancel] + rw [← neg_vsub_eq_vsub_rev p₂ y] + linear_combination (norm := match_scalars <;> field_simp) congr((r₁ + r₂)⁻¹ • $h) + ring rw [lineMap_apply, ← vsub_vadd x p₁, ← vsub_vadd y p₂, vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, ← vadd_assoc, vadd_eq_add, this] exact s.smul_vsub_vadd_mem (r₂ / (r₁ + r₂)) hp₂ hp₁ hp₁ · exact Set.mem_image_of_mem _ - ⟨div_nonneg hr₂.le (Left.add_pos hr₁ hr₂).le, + ⟨by positivity, div_le_one_of_le (le_add_of_nonneg_left hr₁.le) (Left.add_pos hr₁ hr₂).le⟩ theorem SOppSide.exists_sbtw {s : AffineSubspace R P} {x y : P} (h : s.SOppSide x y) : diff --git a/Mathlib/Analysis/Convex/Star.lean b/Mathlib/Analysis/Convex/Star.lean index 44f1a60ccabfb..35d31a5d32be1 100644 --- a/Mathlib/Analysis/Convex/Star.lean +++ b/Mathlib/Analysis/Convex/Star.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Algebra.Order.Group.Instances import Mathlib.Analysis.Convex.Segment import Mathlib.Tactic.GCongr +import Mathlib.Tactic.Module /-! # Star-convex sets @@ -210,14 +211,14 @@ theorem StarConvex.add_left (hs : StarConvex 𝕜 x s) (z : E) : intro y hy a b ha hb hab obtain ⟨y', hy', rfl⟩ := hy refine ⟨a • x + b • y', hs hy' ha hb hab, ?_⟩ - rw [smul_add, smul_add, add_add_add_comm, ← add_smul, hab, one_smul] + match_scalars <;> simp [hab] theorem StarConvex.add_right (hs : StarConvex 𝕜 x s) (z : E) : StarConvex 𝕜 (x + z) ((fun x => x + z) '' s) := by intro y hy a b ha hb hab obtain ⟨y', hy', rfl⟩ := hy refine ⟨a • x + b • y', hs hy' ha hb hab, ?_⟩ - rw [smul_add, smul_add, add_add_add_comm, ← add_smul, hab, one_smul] + match_scalars <;> simp [hab] /-- The translation of a star-convex set is also star-convex. -/ theorem StarConvex.preimage_add_right (hs : StarConvex 𝕜 (z + x) s) : diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 4f746cd1909fd..1027690d2b881 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -46,9 +46,6 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s · positivity · positivity · rw [← add_div, div_self]; positivity - rw [smul_add, smul_add, add_add_add_comm] - nth_rw 2 [add_comm] - rw [← mul_smul, ← mul_smul] classical let w : Fin 3 → 𝕜 := ![az * av * bu, bz * au * bv, au * av] let z : Fin 3 → E := ![p, q, az • x + bz • y] @@ -61,18 +58,15 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s have hw : ∑ i, w i = az * av + bz * au := by trans az * av * bu + (bz * au * bv + au * av) · simp [w, Fin.sum_univ_succ, Fin.sum_univ_zero] - rw [← one_mul (au * av), ← habz, add_mul, ← add_assoc, add_add_add_comm, mul_assoc, ← mul_add, - mul_assoc, ← mul_add, mul_comm av, ← add_mul, ← mul_add, add_comm bu, add_comm bv, habu, - habv, one_mul, mul_one] + linear_combination (au * bv - 1 * au) * habz + (-(1 * az * au) + au) * habv + az * av * habu have hz : ∀ i, z i ∈ ({p, q, az • x + bz • y} : Set E) := fun i => by fin_cases i <;> simp [z] convert (Finset.centerMass_mem_convexHull (Finset.univ : Finset (Fin 3)) (fun i _ => hw₀ i) (by rwa [hw]) fun i _ => hz i : Finset.univ.centerMass w z ∈ _) - rw [Finset.centerMass] - simp_rw [div_eq_inv_mul, hw, mul_assoc, mul_smul (az * av + bz * au)⁻¹, ← smul_add, add_assoc, ← - mul_assoc] + rw [Finset.centerMass, hw] + trans (az * av + bz * au)⁻¹ • + ((az * av * bu) • p + ((bz * au * bv) • q + (au * av) • (az • x + bz • y))) + · module congr 3 - rw [← mul_smul, ← mul_rotate, mul_right_comm, mul_smul, ← mul_smul _ av, mul_rotate, - mul_smul _ bz, ← smul_add] simp only [w, z, smul_add, List.foldr, Matrix.cons_val_succ', Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, add_zero] diff --git a/Mathlib/Analysis/Convex/Strict.lean b/Mathlib/Analysis/Convex/Strict.lean index da771a82ffc55..43a8732b0e6d3 100644 --- a/Mathlib/Analysis/Convex/Strict.lean +++ b/Mathlib/Analysis/Convex/Strict.lean @@ -302,8 +302,7 @@ theorem StrictConvex.eq_of_openSegment_subset_frontier [Nontrivial 𝕜] [Densel theorem StrictConvex.add_smul_mem (hs : StrictConvex 𝕜 s) (hx : x ∈ s) (hxy : x + y ∈ s) (hy : y ≠ 0) {t : 𝕜} (ht₀ : 0 < t) (ht₁ : t < 1) : x + t • y ∈ interior s := by - have h : x + t • y = (1 - t) • x + t • (x + y) := by - rw [smul_add, ← add_assoc, ← _root_.add_smul, sub_add_cancel, one_smul] + have h : x + t • y = (1 - t) • x + t • (x + y) := by match_scalars <;> field_simp rw [h] exact hs hx hxy (fun h => hy <| add_left_cancel (a := x) (by rw [← h, add_zero])) (sub_pos_of_lt ht₁) ht₀ (sub_add_cancel 1 t) @@ -359,16 +358,14 @@ theorem strictConvex_iff_div : StrictConvex 𝕜 s ↔ s.Pairwise fun x y => ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → (a / (a + b)) • x + (b / (a + b)) • y ∈ interior s := - ⟨fun h x hx y hy hxy a b ha hb => by - apply h hx hy hxy (div_pos ha <| add_pos ha hb) (div_pos hb <| add_pos ha hb) - rw [← add_div] - exact div_self (add_pos ha hb).ne', fun h x hx y hy hxy a b ha hb hab => by + ⟨fun h x hx y hy hxy a b ha hb ↦ h hx hy hxy (by positivity) (by positivity) (by field_simp), + fun h x hx y hy hxy a b ha hb hab ↦ by convert h hx hy hxy ha hb <;> rw [hab, div_one]⟩ theorem StrictConvex.mem_smul_of_zero_mem (hs : StrictConvex 𝕜 s) (zero_mem : (0 : E) ∈ s) (hx : x ∈ s) (hx₀ : x ≠ 0) {t : 𝕜} (ht : 1 < t) : x ∈ t • interior s := by - rw [mem_smul_set_iff_inv_smul_mem₀ (zero_lt_one.trans ht).ne'] - exact hs.smul_mem_of_zero_mem zero_mem hx hx₀ (inv_pos.2 <| zero_lt_one.trans ht) (inv_lt_one ht) + rw [mem_smul_set_iff_inv_smul_mem₀ (by positivity)] + exact hs.smul_mem_of_zero_mem zero_mem hx hx₀ (by positivity) (inv_lt_one ht) end AddCommGroup diff --git a/Mathlib/Analysis/NormedSpace/Connected.lean b/Mathlib/Analysis/NormedSpace/Connected.lean index 64d98223dda5e..e14389a41de74 100644 --- a/Mathlib/Analysis/NormedSpace/Connected.lean +++ b/Mathlib/Analysis/NormedSpace/Connected.lean @@ -52,13 +52,11 @@ theorem Set.Countable.isPathConnected_compl_of_one_lt_rank let c := (2 : ℝ)⁻¹ • (a + b) let x := (2 : ℝ)⁻¹ • (b - a) have Ia : c - x = a := by - simp only [c, x, smul_add, smul_sub] - abel_nf - simp [← Int.cast_smul_eq_zsmul ℝ 2] + simp only [c, x] + module have Ib : c + x = b := by - simp only [c, x, smul_add, smul_sub] - abel_nf - simp [← Int.cast_smul_eq_zsmul ℝ 2] + simp only [c, x] + module have x_ne_zero : x ≠ 0 := by simpa [x] using sub_ne_zero.2 hab.symm obtain ⟨y, hy⟩ : ∃ y, LinearIndependent ℝ ![x, y] := exists_linearIndependent_pair_of_one_lt_rank h x_ne_zero diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean index b933762e6b4a6..7f40cc129de8e 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean @@ -65,9 +65,8 @@ def rotation (θ : Real.Angle) : V ≃ₗᵢ[ℝ] V := · simp only [o.rightAngleRotation_rightAngleRotation, o.rotationAux_apply, Function.comp_apply, id, LinearEquiv.coe_coe, LinearIsometry.coe_toLinearMap, LinearIsometryEquiv.coe_toLinearEquiv, map_smul, map_sub, LinearMap.coe_comp, - LinearMap.id_coe, LinearMap.smul_apply, LinearMap.sub_apply, ← mul_smul, add_smul, - smul_add, smul_neg, smul_sub, mul_comm, sq] - abel + LinearMap.id_coe, LinearMap.smul_apply, LinearMap.sub_apply] + module · simp) (by ext x @@ -75,10 +74,8 @@ def rotation (θ : Real.Angle) : V ≃ₗᵢ[ℝ] V := · simp only [o.rightAngleRotation_rightAngleRotation, o.rotationAux_apply, Function.comp_apply, id, LinearEquiv.coe_coe, LinearIsometry.coe_toLinearMap, LinearIsometryEquiv.coe_toLinearEquiv, map_add, map_smul, LinearMap.coe_comp, - LinearMap.id_coe, LinearMap.smul_apply, LinearMap.sub_apply, - add_smul, smul_neg, smul_sub, smul_smul] - ring_nf - abel + LinearMap.id_coe, LinearMap.smul_apply, LinearMap.sub_apply] + module · simp) theorem rotation_apply (θ : Real.Angle) (x : V) : @@ -146,11 +143,9 @@ theorem rotation_pi_div_two : o.rotation (π / 2 : ℝ) = J := by @[simp] theorem rotation_rotation (θ₁ θ₂ : Real.Angle) (x : V) : o.rotation θ₁ (o.rotation θ₂ x) = o.rotation (θ₁ + θ₂) x := by - simp only [o.rotation_apply, ← mul_smul, Real.Angle.cos_add, Real.Angle.sin_add, add_smul, - sub_smul, LinearIsometryEquiv.trans_apply, smul_add, LinearIsometryEquiv.map_add, - LinearIsometryEquiv.map_smul, rightAngleRotation_rightAngleRotation, smul_neg] - ring_nf - abel + simp only [o.rotation_apply, Real.Angle.cos_add, Real.Angle.sin_add, LinearIsometryEquiv.map_add, + LinearIsometryEquiv.trans_apply, map_smul, rightAngleRotation_rightAngleRotation] + module /-- Rotating twice is equivalent to rotating by the sum of the angles. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index ed68d3b7b7e44..f045983a13cd3 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -11,7 +11,8 @@ import Mathlib.Tactic.FinCases import Mathlib.Tactic.LinearCombination import Mathlib.Lean.Expr.ExtraRecognizers import Mathlib.Data.Set.Subsingleton -import Mathlib.Tactic.Abel +import Mathlib.Tactic.Module +import Mathlib.Tactic.NoncommRing /-! @@ -596,9 +597,8 @@ theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := by have : (s - s') • x + (t - t') • y = 0 := by - rw [← sub_eq_zero_of_eq h', ← sub_eq_zero] - simp only [sub_smul] - abel + rw [← sub_eq_zero_of_eq h'] + match_scalars <;> noncomm_ring simpa [sub_eq_zero] using h.eq_zero_of_pair this lemma LinearIndependent.eq_zero_of_pair' {x y : M} (h : LinearIndependent R ![x, y]) @@ -616,8 +616,7 @@ lemma LinearIndependent.linear_combination_pair_of_det_ne_zero {R M : Type*} [Co apply LinearIndependent.pair_iff.2 (fun s t hst ↦ ?_) have H : (s * a + t * c) • x + (s * b + t * d) • y = 0 := by convert hst using 1 - simp only [_root_.add_smul, smul_add, smul_smul] - abel + module have I1 : s * a + t * c = 0 := (h.eq_zero_of_pair H).1 have I2 : s * b + t * d = 0 := (h.eq_zero_of_pair H).2 have J1 : (a * d - b * c) * s = 0 := by linear_combination d * I1 - c * I2 @@ -1111,11 +1110,10 @@ theorem linearIndependent_monoidHom (G : Type*) [Monoid G] (L : Type*) [CommRing rw [Finset.sum_insert has, Finset.sum_insert has] _ = (∑ i ∈ insert a s, g i * i (x * y)) - - ∑ i ∈ insert a s, a x * (g i * i y) := - congr - (congr_arg Sub.sub - (Finset.sum_congr rfl fun i _ => by rw [i.map_mul, mul_assoc])) - (Finset.sum_congr rfl fun _ _ => by rw [mul_assoc, mul_left_comm]) + ∑ i ∈ insert a s, a x * (g i * i y) := by + congrm ∑ i ∈ insert a s, ?_ - ∑ i ∈ insert a s, ?_ + · rw [map_mul, mul_assoc] + · rw [mul_assoc, mul_left_comm] _ = (∑ i ∈ insert a s, (g i • (i : G → L))) (x * y) - a x * (∑ i ∈ insert a s, (g i • (i : G → L))) y := by @@ -1231,7 +1229,7 @@ theorem mem_span_insert_exchange : have a0 : a ≠ 0 := by rintro rfl simp_all - simp [a0, smul_add, smul_smul] + match_scalars <;> simp [a0] theorem linearIndependent_iff_not_mem_span : LinearIndependent K v ↔ ∀ i, v i ∉ span K (v '' (univ \ {i})) := by @@ -1305,8 +1303,8 @@ theorem LinearIndependent.pair_iff' {x y : V} (hx : x ≠ 0) : by_cases ht : t = 0 · exact ⟨by simpa [ht, hx] using hst, ht⟩ apply_fun (t⁻¹ • ·) at hst - simp only [smul_add, smul_smul, inv_mul_cancel₀ ht, one_smul, smul_zero] at hst - cases H (-(t⁻¹ * s)) (by rwa [neg_smul, neg_eq_iff_eq_neg, eq_neg_iff_add_eq_zero]) + simp only [smul_add, smul_smul, inv_mul_cancel₀ ht] at hst + cases H (-(t⁻¹ * s)) <| by linear_combination (norm := match_scalars <;> noncomm_ring) -hst theorem linearIndependent_fin_cons {n} {v : Fin n → V} : LinearIndependent K (Fin.cons x v : Fin (n + 1) → V) ↔ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index caedda89e57e7..7eba713353773 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -854,15 +854,15 @@ theorem associated_toQuadraticMap (B : BilinMap R M R) (x y : M) : theorem associated_left_inverse (h : B₁.IsSymm) : associatedHom S B₁.toQuadraticMap = B₁ := LinearMap.ext₂ fun x y => by - rw [associated_toQuadraticMap, ← h.eq x y, RingHom.id_apply, ← two_mul, ← smul_mul_assoc, - smul_eq_mul, invOf_mul_self, one_mul] + rw [associated_toQuadraticMap, ← h.eq x y, RingHom.id_apply] + match_scalars + linear_combination invOf_mul_self' (2:R) -- Porting note: moved from below to golf the next theorem theorem associated_eq_self_apply (x : M) : associatedHom S Q x x = Q x := by - rw [associated_apply, map_add_self, ← three_add_one_eq_four, ← two_add_one_eq_three, add_smul, - add_smul, one_smul, add_sub_cancel_right, add_sub_cancel_right, two_smul, ← two_smul R, - ← smul_assoc] - simp only [smul_eq_mul, invOf_mul_self', one_smul] + rw [associated_apply, map_add_self] + match_scalars + linear_combination invOf_mul_self' (2:R) theorem toQuadraticMap_associated : (associatedHom S Q).toQuadraticMap = Q := QuadraticMap.ext <| associated_eq_self_apply S Q @@ -1288,8 +1288,7 @@ theorem basisRepr_eq_of_iIsOrtho {R M} [CommRing R] [AddCommGroup M] [Module R M smul_eq_mul, smul_eq_mul] ring_nf · intro i _ hij - rw [LinearMap.map_smul, LinearMap.map_smul₂, - show associatedHom R Q (v i) (v j) = 0 from hv₂ hij, smul_eq_mul, smul_eq_mul, - mul_zero, mul_zero] + rw [LinearMap.map_smul, LinearMap.map_smul₂, hv₂ hij] + module end QuadraticMap diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index 87778718498eb..547a7c8540f4b 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -6,6 +6,7 @@ Authors: Joseph Myers import Mathlib.Algebra.Order.Module.Algebra import Mathlib.LinearAlgebra.LinearIndependent import Mathlib.Algebra.Ring.Subring.Units +import Mathlib.Tactic.Positivity /-! # Rays in modules @@ -106,7 +107,7 @@ lemma sameRay_nonneg_smul_right (v : M) (h : 0 ≤ a) : SameRay R v (a • v) := · rw [← algebraMap_smul R a v, h, zero_smul] exact zero_right _ · refine Or.inr <| Or.inr ⟨algebraMap S R a, 1, h, by nontriviality R; exact zero_lt_one, ?_⟩ - rw [algebraMap_smul, one_smul] + module /-- A nonnegative multiple of a vector is in the same ray as that vector. -/ lemma sameRay_nonneg_smul_left (v : M) (ha : 0 ≤ a) : SameRay R (a • v) v := @@ -170,9 +171,8 @@ theorem add_left (hx : SameRay R x z) (hy : SameRay R y z) : SameRay R (x + y) z rcases hx.exists_pos hx₀ hz₀ with ⟨rx, rz₁, hrx, hrz₁, Hx⟩ rcases hy.exists_pos hy₀ hz₀ with ⟨ry, rz₂, hry, hrz₂, Hy⟩ refine Or.inr (Or.inr ⟨rx * ry, ry * rz₁ + rx * rz₂, mul_pos hrx hry, ?_, ?_⟩) - · apply_rules [add_pos, mul_pos] - · simp only [mul_smul, smul_add, add_smul, ← Hx, ← Hy] - rw [smul_comm] + · positivity + · convert congr(ry • $Hx + rx • $Hy) using 1 <;> module /-- If `y` and `z` are on the same ray as `x`, then so is `y + z`. -/ theorem add_right (hy : SameRay R x y) (hz : SameRay R x z) : SameRay R x (y + z) := @@ -531,11 +531,11 @@ theorem sameRay_or_sameRay_neg_iff_not_linearIndependent {x y : M} : rcases lt_trichotomy (m 1) 0 with (hm1 | hm1 | hm1) · refine Or.inr (Or.inr (Or.inr ⟨-m 0, -m 1, Left.neg_pos_iff.2 hm0, Left.neg_pos_iff.2 hm1, ?_⟩)) - rw [neg_smul_neg, neg_smul, hm, neg_neg] + linear_combination (norm := module) -hm · exfalso simp [hm1, hx, hm0.ne] at hm · refine Or.inl (Or.inr (Or.inr ⟨-m 0, m 1, Left.neg_pos_iff.2 hm0, hm1, ?_⟩)) - rw [neg_smul, hm, neg_neg] + linear_combination (norm := module) -hm · exfalso simp [hm0, hy, hm1.ne] at hm · rw [Fin.exists_fin_two] at hmne diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index df6dd1f46c15f..6c8bfa2572edc 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -427,10 +427,8 @@ theorem functional_equation₀ (s : ℂ) : P.Λ₀ (P.k - s) = P.ε • P.symm. /-- Functional equation formulated for `Λ`. -/ theorem functional_equation (s : ℂ) : P.Λ (P.k - s) = P.ε • P.symm.Λ s := by - have := P.functional_equation₀ s - rw [P.Λ₀_eq, P.symm_Λ₀_eq, sub_sub_cancel] at this - rwa [smul_add, smul_add, ← mul_smul, mul_one_div, ← mul_smul, ← mul_div_assoc, - mul_inv_cancel₀ P.hε, add_assoc, add_comm (_ • _), add_assoc, add_left_inj] at this + linear_combination (norm := module) P.functional_equation₀ s - P.Λ₀_eq (P.k - s) + + congr(P.ε • $(P.symm_Λ₀_eq s)) + congr(($(mul_inv_cancel₀ P.hε) / ((P.k:ℂ) - s)) • P.f₀) /-- The residue of `Λ` at `s = k` is equal to `ε • g₀`. -/ theorem Λ_residue_k : @@ -444,8 +442,7 @@ theorem Λ_residue_k : exact continuousAt_const.div continuousAt_id (ofReal_ne_zero.mpr P.hk.ne') · refine (tendsto_const_nhds.mono_left nhdsWithin_le_nhds).congr' ?_ refine eventually_nhdsWithin_of_forall (fun s (hs : s ≠ P.k) ↦ ?_) - simp_rw [← mul_smul] - congr 1 + match_scalars field_simp [sub_ne_zero.mpr hs.symm] ring @@ -457,7 +454,7 @@ theorem Λ_residue_zero : · exact (continuous_id.smul P.differentiable_Λ₀.continuous).tendsto _ · refine (tendsto_const_nhds.mono_left nhdsWithin_le_nhds).congr' ?_ refine eventually_nhdsWithin_of_forall (fun s (hs : s ≠ 0) ↦ ?_) - simp_rw [← mul_smul] + match_scalars field_simp [sub_ne_zero.mpr hs.symm] · rw [show 𝓝 0 = 𝓝 ((0 : ℂ) • (P.ε / (P.k - 0 : ℂ)) • P.g₀) by rw [zero_smul]] exact (continuousAt_id.smul ((continuousAt_const.div ((continuous_sub_left _).continuousAt) From e58fb43f3f5b3d77ecb7eb6d351b78ee632435a2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 3 Oct 2024 16:01:26 +0000 Subject: [PATCH 155/174] feat(Topology/Group): drop an unneeded assumption (#16551) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - prove that the quotient map `G → G ⧸ N` is an open quotient map; - use recently added lemmas to drop a `[LocallyCompactSpace G]` assumption in `QuotientGroup.continuousSMul` (now called `QuotientGroup.instContinuousSMul`). --- .../MeasureTheory/Measure/Haar/Quotient.lean | 2 +- Mathlib/Topology/Algebra/ConstMulAction.lean | 5 ++ Mathlib/Topology/Algebra/Group/Basic.lean | 57 ++++++++++++------- Mathlib/Topology/Algebra/Group/Compact.lean | 32 ++--------- Mathlib/Topology/Algebra/Module/Basic.lean | 19 +++---- Mathlib/Topology/Algebra/ProperAction.lean | 8 +-- Mathlib/Topology/Algebra/Ring/Ideal.lean | 30 ++++------ Mathlib/Topology/Maps/OpenQuotient.lean | 6 +- 8 files changed, 74 insertions(+), 85 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index 4882bfa945404..6cb8d92d7d95f 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -223,7 +223,7 @@ theorem MeasureTheory.QuotientMeasureEqMeasurePreimage.haarMeasure_quotient [Loc [IsFiniteMeasure μ] : IsHaarMeasure μ := by obtain ⟨K⟩ := PositiveCompacts.nonempty' (α := G) let K' : PositiveCompacts (G ⧸ Γ) := - K.map π continuous_coinduced_rng (QuotientGroup.isOpenMap_coe Γ) + K.map π QuotientGroup.continuous_mk QuotientGroup.isOpenMap_coe haveI : IsMulLeftInvariant μ := MeasureTheory.QuotientMeasureEqMeasurePreimage.mulInvariantMeasure_quotient ν rw [haarMeasure_unique μ K'] diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 482927bb1d468..39bd780335969 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -451,6 +451,11 @@ theorem isOpenMap_quotient_mk'_mul [ContinuousConstSMul Γ T] : rw [isOpen_coinduced, MulAction.quotient_preimage_image_eq_union_mul U] exact isOpen_iUnion fun γ => isOpenMap_smul γ U hU +@[to_additive] +theorem MulAction.isOpenQuotientMap_quotientMk [ContinuousConstSMul Γ T] : + IsOpenQuotientMap (Quotient.mk (MulAction.orbitRel Γ T)) := + ⟨surjective_quot_mk _, continuous_quot_mk, isOpenMap_quotient_mk'_mul⟩ + /-- The quotient by a discontinuous group action of a locally compact t2 space is t2. -/ @[to_additive "The quotient by a discontinuous group action of a locally compact t2 space is t2."] diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 056e62ee57db1..99386fcf912e5 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -7,6 +7,7 @@ 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.Order.Archimedean.Basic import Mathlib.GroupTheory.QuotientGroup.Basic @@ -862,16 +863,24 @@ instance [CompactSpace G] (N : Subgroup G) : CompactSpace (G ⧸ N) := theorem quotientMap_mk (N : Subgroup G) : QuotientMap (mk : G → G ⧸ N) := quotientMap_quot_mk -variable [TopologicalGroup G] (N : Subgroup G) +@[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 isOpenMap_coe : IsOpenMap ((↑) : G → G ⧸ N) := - isOpenMap_quotient_mk'_mul +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 := - letI := leftRel N -- `Dense.quotient` assumes `[Setoid G]` - ⟨fun h ↦ h.quotient.mono <| image_preimage_subset _ _, fun h ↦ h.preimage <| isOpenMap_coe _⟩ + isOpenQuotientMap_mk.dense_preimage_iff @[to_additive] theorem dense_image_mk {s : Set G} : @@ -879,27 +888,18 @@ theorem dense_image_mk {s : Set G} : rw [← dense_preimage_mk, preimage_image_mk_eq_mul] @[to_additive] -instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where - continuous_mul := by - have cont : Continuous (((↑) : G → G ⧸ N) ∘ fun p : G × G ↦ p.fst * p.snd) := - continuous_quot_mk.comp continuous_mul - have quot : QuotientMap fun p : G × G ↦ ((p.1 : G ⧸ N), (p.2 : G ⧸ N)) := by - apply IsOpenMap.to_quotientMap - · exact (QuotientGroup.isOpenMap_coe N).prod (QuotientGroup.isOpenMap_coe N) - · exact continuous_quot_mk.prod_map continuous_quot_mk - · exact (surjective_quot_mk _).prodMap (surjective_quot_mk _) - exact quot.continuous_iff.2 cont - continuous_inv := continuous_inv.quotient_map' _ +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 (attr := deprecated (since := "2024-08-05"))] -theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := - instTopologicalGroup N +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) := - le_antisymm ((QuotientGroup.isOpenMap_coe N).nhds_le x) continuous_quot_mk.continuousAt + (isOpenQuotientMap_mk.map_nhds_eq _).symm @[to_additive] instance instFirstCountableTopology [FirstCountableTopology G] : @@ -911,6 +911,21 @@ 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 @@ -1612,7 +1627,7 @@ instance [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N 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 N).image_mem_nhds s_mem, mapsTo'.mp hs, + exact ⟨π '' s, QuotientGroup.isOpenMap_coe.image_mem_nhds s_mem, mapsTo'.mp hs, s_comp.image C⟩ end diff --git a/Mathlib/Topology/Algebra/Group/Compact.lean b/Mathlib/Topology/Algebra/Group/Compact.lean index 57b0155e2e6a8..b479c867ca91f 100644 --- a/Mathlib/Topology/Algebra/Group/Compact.lean +++ b/Mathlib/Topology/Algebra/Group/Compact.lean @@ -4,24 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Topology.Algebra.Group.Basic -import Mathlib.Topology.CompactOpen import Mathlib.Topology.Sets.Compacts /-! # Additional results on topological groups -Two results on topological groups that have been separated out as they require more substantial -imports developing either positive compacts or the compact open topology. - +A result on topological groups that has been separated out +as it requires more substantial imports developing positive compacts. -/ -universe u v w x - -variable {α : Type u} {β : Type v} {G : Type w} {H : Type x} -section - -variable [TopologicalSpace G] [Group G] [TopologicalGroup G] +universe u +variable {G : Type u} [TopologicalSpace G] [Group G] [TopologicalGroup G] /-- Every topological group in which there exists a compact set with nonempty interior is locally compact. -/ @@ -32,21 +26,3 @@ theorem TopologicalSpace.PositiveCompacts.locallyCompactSpace_of_group (K : PositiveCompacts G) : LocallyCompactSpace G := let ⟨_x, hx⟩ := K.interior_nonempty K.isCompact.locallyCompactSpace_of_mem_nhds_of_group (mem_interior_iff_mem_nhds.1 hx) - -end - -section Quotient - -variable [Group G] [TopologicalSpace G] [TopologicalGroup G] {Γ : Subgroup G} - -@[to_additive] -instance QuotientGroup.continuousSMul [LocallyCompactSpace G] : ContinuousSMul G (G ⧸ Γ) where - continuous_smul := by - let F : G × G ⧸ Γ → G ⧸ Γ := fun p => p.1 • p.2 - change Continuous F - have H : Continuous (F ∘ fun p : G × G => (p.1, QuotientGroup.mk p.2)) := by - change Continuous fun p : G × G => QuotientGroup.mk (p.1 * p.2) - exact continuous_coinduced_rng.comp continuous_mul - exact QuotientMap.continuous_lift_prod_right quotientMap_quotient_mk' H - -end Quotient diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 36ed9e6021dea..99adb17d2c59c 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -2375,21 +2375,20 @@ variable {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] [TopologicalSpace instance _root_.QuotientModule.Quotient.topologicalSpace : TopologicalSpace (M ⧸ S) := inferInstanceAs (TopologicalSpace (Quotient S.quotientRel)) -theorem isOpenMap_mkQ [TopologicalAddGroup M] : IsOpenMap S.mkQ := - QuotientAddGroup.isOpenMap_coe S.toAddSubgroup +theorem isOpenMap_mkQ [ContinuousAdd M] : IsOpenMap S.mkQ := + QuotientAddGroup.isOpenMap_coe + +theorem isOpenQuotientMap_mkQ [ContinuousAdd M] : IsOpenQuotientMap S.mkQ := + QuotientAddGroup.isOpenQuotientMap_mk instance topologicalAddGroup_quotient [TopologicalAddGroup M] : TopologicalAddGroup (M ⧸ S) := inferInstanceAs <| TopologicalAddGroup (M ⧸ S.toAddSubgroup) instance continuousSMul_quotient [TopologicalSpace R] [TopologicalAddGroup M] [ContinuousSMul R M] : - ContinuousSMul R (M ⧸ S) := by - constructor - have quot : QuotientMap fun au : R × M => (au.1, S.mkQ au.2) := - IsOpenMap.to_quotientMap (IsOpenMap.id.prod S.isOpenMap_mkQ) - (continuous_id.prod_map continuous_quot_mk) - (Function.surjective_id.prodMap <| surjective_quot_mk _) - rw [quot.continuous_iff] - exact continuous_quot_mk.comp continuous_smul + ContinuousSMul R (M ⧸ S) where + continuous_smul := by + rw [← (IsOpenQuotientMap.id.prodMap S.isOpenQuotientMap_mkQ).continuous_comp_iff] + exact continuous_quot_mk.comp continuous_smul instance t3_quotient_of_isClosed [TopologicalAddGroup M] [IsClosed (S : Set M)] : T3Space (M ⧸ S) := diff --git a/Mathlib/Topology/Algebra/ProperAction.lean b/Mathlib/Topology/Algebra/ProperAction.lean index cd13f848a2f41..b937385d62d95 100644 --- a/Mathlib/Topology/Algebra/ProperAction.lean +++ b/Mathlib/Topology/Algebra/ProperAction.lean @@ -129,11 +129,9 @@ theorem t2Space_quotient_mulAction_of_properSMul [ProperSMul G X] : rw [t2_iff_isClosed_diagonal] set R := MulAction.orbitRel G X let π : X → Quotient R := Quotient.mk' - have : QuotientMap (Prod.map π π) := - (isOpenMap_quotient_mk'_mul.prod isOpenMap_quotient_mk'_mul).to_quotientMap - (continuous_quotient_mk'.prod_map continuous_quotient_mk') - ((surjective_quotient_mk' _).prodMap (surjective_quotient_mk' _)) - rw [← this.isClosed_preimage] + have : IsOpenQuotientMap (Prod.map π π) := + MulAction.isOpenQuotientMap_quotientMk.prodMap MulAction.isOpenQuotientMap_quotientMk + rw [← this.quotientMap.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 bd45ee761f821..aa3cf06eb57a1 100644 --- a/Mathlib/Topology/Algebra/Ring/Ideal.lean +++ b/Mathlib/Topology/Algebra/Ring/Ideal.lean @@ -51,26 +51,18 @@ instance topologicalRingQuotientTopology : TopologicalSpace (R ⧸ N) := -- note for the reader: in the following, `mk` is `Ideal.Quotient.mk`, the canonical map `R → R/I`. variable [TopologicalRing R] -theorem QuotientRing.isOpenMap_coe : IsOpenMap (mk N) := by - intro s s_op - change IsOpen (mk N ⁻¹' (mk N '' s)) - rw [quotient_ring_saturate] - exact isOpen_iUnion fun ⟨n, _⟩ => isOpenMap_add_left n s s_op +theorem QuotientRing.isOpenMap_coe : IsOpenMap (mk N) := + QuotientAddGroup.isOpenMap_coe + +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) := - IsOpenMap.to_quotientMap ((QuotientRing.isOpenMap_coe N).prod (QuotientRing.isOpenMap_coe N)) - ((continuous_quot_mk.comp continuous_fst).prod_mk (continuous_quot_mk.comp continuous_snd)) - (by rintro ⟨⟨x⟩, ⟨y⟩⟩; exact ⟨(x, y), rfl⟩) - -instance topologicalRing_quotient : TopologicalRing (R ⧸ N) := - TopologicalSemiring.toTopologicalRing - { continuous_add := - have cont : Continuous (mk N ∘ fun p : R × R => p.fst + p.snd) := - continuous_quot_mk.comp continuous_add - (QuotientMap.continuous_iff (QuotientRing.quotientMap_coe_coe N)).mpr cont - continuous_mul := - have cont : Continuous (mk N ∘ fun p : R × R => p.fst * p.snd) := - continuous_quot_mk.comp continuous_mul - (QuotientMap.continuous_iff (QuotientRing.quotientMap_coe_coe N)).mpr cont } + ((isOpenQuotientMap_mk N).prodMap (isOpenQuotientMap_mk N)).quotientMap + +instance topologicalRing_quotient : TopologicalRing (R ⧸ N) where + __ := QuotientAddGroup.instTopologicalAddGroup _ + continuous_mul := (QuotientRing.quotientMap_coe_coe N).continuous_iff.2 <| + continuous_quot_mk.comp continuous_mul end CommRing diff --git a/Mathlib/Topology/Maps/OpenQuotient.lean b/Mathlib/Topology/Maps/OpenQuotient.lean index f5dc76959e06b..88433f7a190fd 100644 --- a/Mathlib/Topology/Maps/OpenQuotient.lean +++ b/Mathlib/Topology/Maps/OpenQuotient.lean @@ -22,7 +22,7 @@ Contrary to general quotient maps, the category of open quotient maps is closed under `Prod.map`. -/ -open Function Filter +open Function Set Filter open scoped Topology variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : X → Y} @@ -57,4 +57,8 @@ theorem continuousAt_comp_iff (h : IsOpenQuotientMap f) {g : Y → Z} {x : X} : ContinuousAt (g ∘ f) x ↔ ContinuousAt g (f x) := by simp only [ContinuousAt, ← h.map_nhds_eq, tendsto_map'_iff, comp_def] +theorem dense_preimage_iff (h : IsOpenQuotientMap f) {s : Set Y} : Dense (f ⁻¹' s) ↔ Dense s := + ⟨fun hs ↦ h.surjective.denseRange.dense_of_mapsTo h.continuous hs (mapsTo_preimage _ _), + fun hs ↦ hs.preimage h.isOpenMap⟩ + end IsOpenQuotientMap From 1a4c6bcc621d0b5f485d6121d26937525a458c54 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 3 Oct 2024 17:24:54 +0000 Subject: [PATCH 156/174] feat(AlgebraicGeometry): Residue fields of schemes. (#15333) --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/ResidueField.lean | 180 ++++++++++++++++++++ 2 files changed, 181 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/ResidueField.lean diff --git a/Mathlib.lean b/Mathlib.lean index 22ea50c9663f9..038ecf2264b60 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -874,6 +874,7 @@ import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Topology import Mathlib.AlgebraicGeometry.Properties import Mathlib.AlgebraicGeometry.Pullbacks +import Mathlib.AlgebraicGeometry.ResidueField import Mathlib.AlgebraicGeometry.Restrict import Mathlib.AlgebraicGeometry.Scheme import Mathlib.AlgebraicGeometry.Sites.BigZariski diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean new file mode 100644 index 0000000000000..f192506ed485b --- /dev/null +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -0,0 +1,180 @@ +/- +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 + +/-! + +# Residue fields of points + +## Main definitions + +The following are in the `AlgebraicGeometry.Scheme` namespace: + +- `AlgebraicGeometry.Scheme.residueField`: The residue field of the stalk at `x`. +- `AlgebraicGeometry.Scheme.evaluation`: For open subsets `U` of `X` containing `x`, + the evaluation map from sections over `U` to the residue field at `x`. +- `AlgebraicGeometry.Scheme.Hom.residueFieldMap`: A morphism of schemes induce a homomorphism of + residue fields. +- `AlgebraicGeometry.Scheme.fromSpecResidueField`: The canonical map `Spec κ(x) ⟶ X`. + +-/ + +universe u + +open CategoryTheory TopologicalSpace Opposite + +noncomputable section + +namespace AlgebraicGeometry.Scheme + +variable (X : Scheme.{u}) {U : X.Opens} + +/-- The residue field of `X` at a point `x` is the residue field of the stalk of `X` +at `x`. -/ +def residueField (x : X) : CommRingCat := + CommRingCat.of <| LocalRing.ResidueField (X.presheaf.stalk x) + +instance (x : X) : Field (X.residueField x) := + inferInstanceAs <| Field (LocalRing.ResidueField (X.presheaf.stalk x)) + +/-- The residue map from the stalk to the residue field. -/ +def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := + LocalRing.residue _ + +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 `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`. + +If we interpret sections over `U` as functions of `X` defined on `U`, then this ring map +corresponds to evaluation at `x`. +-/ +def evaluation (U : X.Opens) (x : X) (hx : x ∈ U) : Γ(X, U) ⟶ X.residueField x := + X.presheaf.germ ⟨x, hx⟩ ≫ X.residue _ + +@[reassoc] +lemma germ_residue (x : U) : X.presheaf.germ x ≫ X.residue x.1 = X.evaluation U x x.2 := rfl + +/-- The global evaluation map from `Γ(X, ⊤)` to the residue field at `x`. -/ +abbrev Γevaluation (x : X) : Γ(X, ⊤) ⟶ X.residueField x := + X.evaluation ⊤ x trivial + +@[simp] +lemma evaluation_eq_zero_iff_not_mem_basicOpen (x : X) (hx : x ∈ U) (f : Γ(X, U)) : + X.evaluation U x hx f = 0 ↔ x ∉ X.basicOpen f := + X.toLocallyRingedSpace.evaluation_eq_zero_iff_not_mem_basicOpen ⟨x, hx⟩ f + +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 + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) + +instance (x) : IsLocalRingHom (f.stalkMap x) := inferInstanceAs (IsLocalRingHom (f.val.stalkMap 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 := + 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 + simp [Hom.residueFieldMap] + rfl + +@[simp] +lemma residueFieldMap_id (x : X) : + Hom.residueFieldMap (𝟙 X) x = 𝟙 (X.residueField x) := + LocallyRingedSpace.residueFieldMap_id _ + +@[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 := + 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 = + f.app V ≫ X.evaluation (f ⁻¹ᵁ V) x hx := + LocallyRingedSpace.evaluation_naturality f ⟨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) = + X.evaluation (f ⁻¹ᵁ V) x hx (f.app V s) := + LocallyRingedSpace.evaluation_naturality_apply f ⟨x, hx⟩ s + +instance [IsOpenImmersion f] (x) : IsIso (f.residueFieldMap x) := + (LocalRing.ResidueField.mapEquiv + (asIso (f.stalkMap x)).commRingCatIsoToRingEquiv).toCommRingCatIso.isIso_hom + +section congr + +-- replace this def if hard to work with +/-- The isomorphism between residue fields of equal points. -/ +def residueFieldCongr {x y : X} (h : x = y) : + X.residueField x ≅ X.residueField y := + eqToIso (by subst h; rfl) + +@[simp] +lemma residueFieldCongr_refl {x : X} : + X.residueFieldCongr (refl x) = Iso.refl _ := rfl + +@[simp] +lemma residueFieldCongr_symm {x y : X} (e : x = y) : + (X.residueFieldCongr e).symm = X.residueFieldCongr e.symm := rfl + +@[simp] +lemma residueFieldCongr_inv {x y : X} (e : x = y) : + (X.residueFieldCongr e).inv = (X.residueFieldCongr e.symm).hom := rfl + +@[simp] +lemma residueFieldCongr_trans {x y z : X} (e : x = y) (e' : y = z) : + X.residueFieldCongr e ≪≫ X.residueFieldCongr e' = X.residueFieldCongr (e.trans e') := by + subst e e' + rfl + +@[reassoc (attr := simp)] +lemma residueFieldCongr_trans_hom (X : Scheme) {x y z : X} (e : x = y) (e' : y = z) : + (X.residueFieldCongr e).hom ≫ (X.residueFieldCongr e').hom = + (X.residueFieldCongr (e.trans e')).hom := by + subst e e' + rfl + +@[reassoc] +lemma residue_residueFieldCongr (X : Scheme) {x y : X} (h : x = y) : + X.residue x ≫ (X.residueFieldCongr h).hom = + (X.presheaf.stalkCongr (.of_eq h)).hom ≫ X.residue y := by + subst h + simp + +end congr + +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 + +@[reassoc (attr := simp)] +lemma residueFieldCongr_fromSpecResidueField {x y : X} (h : x = y) : + Spec.map (X.residueFieldCongr h).hom ≫ X.fromSpecResidueField _ = + X.fromSpecResidueField _ := by + subst h; simp + +end fromResidueField + +end Scheme + +end AlgebraicGeometry From 6b2a12273e5ba90af2d475ff0d13a2e37e87cddc Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 3 Oct 2024 17:34:25 +0000 Subject: [PATCH 157/174] feat(RingTheory/Unramified): Classification of unramifield field extensions. (#15123) Co-authored-by: Yury G. Kudryashov --- Mathlib.lean | 1 + Mathlib/Algebra/Polynomial/Taylor.lean | 7 + Mathlib/RingTheory/Unramified/Field.lean | 213 +++++++++++++++++++++++ 3 files changed, 221 insertions(+) create mode 100644 Mathlib/RingTheory/Unramified/Field.lean diff --git a/Mathlib.lean b/Mathlib.lean index 038ecf2264b60..b1d7ca948d1d3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4089,6 +4089,7 @@ 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.Valuation.AlgebraInstances import Mathlib.RingTheory.Valuation.Basic diff --git a/Mathlib/Algebra/Polynomial/Taylor.lean b/Mathlib/Algebra/Polynomial/Taylor.lean index 6f9f0e82544ef..4a89d16c0bea7 100644 --- a/Mathlib/Algebra/Polynomial/Taylor.lean +++ b/Mathlib/Algebra/Polynomial/Taylor.lean @@ -125,4 +125,11 @@ theorem sum_taylor_eq {R} [CommRing R] (f : R[X]) (r : R) : rw [← comp_eq_sum_left, sub_eq_add_neg, ← C_neg, ← taylor_apply, taylor_taylor, neg_add_cancel, taylor_zero] +theorem eval_add_of_sq_eq_zero {A} [CommSemiring A] (p : Polynomial A) (x y : A) (hy : y ^ 2 = 0) : + p.eval (x + y) = p.eval x + p.derivative.eval x * y := by + rw [add_comm, ← Polynomial.taylor_eval, + Polynomial.eval_eq_sum_range' ((Nat.lt_succ_self _).trans (Nat.lt_succ_self _)), + Finset.sum_range_succ', Finset.sum_range_succ'] + simp [pow_succ, mul_assoc, ← pow_two, hy, add_comm (eval x p)] + end Polynomial diff --git a/Mathlib/RingTheory/Unramified/Field.lean b/Mathlib/RingTheory/Unramified/Field.lean new file mode 100644 index 0000000000000..3a0152d738c1c --- /dev/null +++ b/Mathlib/RingTheory/Unramified/Field.lean @@ -0,0 +1,213 @@ +/- +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.FieldTheory.PurelyInseparable +import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.Algebra.Polynomial.Taylor +import Mathlib.RingTheory.Unramified.Finite + +/-! +# Unramified algebras over fields + +## Main results + +Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. + +- `Algebra.FormallyUnramified.bijective_of_isAlgClosed_of_localRing`: + If `A` is `K`-unramified and `K` is alg-closed, then `K = A`. +- `Algebra.FormallyUnramified.isReduced_of_field`: + If `A` is `K`-unramified then `A` is reduced. +- `Algebra.FormallyUnramified.iff_isSeparable`: + `L` is unramified over `K` iff `L` is separable over `K`. + +## 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.FormallyUnramified + +theorem of_isSeparable [Algebra.IsSeparable K L] : FormallyUnramified K L := by + constructor + intros B _ _ I hI f₁ f₂ e + ext x + have : f₁ x - f₂ x ∈ I := by + simpa [Ideal.Quotient.mk_eq_mk_iff_sub_mem] using AlgHom.congr_fun e x + have := Polynomial.eval_add_of_sq_eq_zero ((minpoly K x).map (algebraMap K B)) (f₂ x) + (f₁ x - f₂ x) (show (f₁ x - f₂ x) ^ 2 ∈ ⊥ from hI ▸ Ideal.pow_mem_pow this 2) + simp only [add_sub_cancel, eval_map_algebraMap, aeval_algHom_apply, minpoly.aeval, map_zero, + derivative_map, zero_add] at this + rwa [eq_comm, ((isUnit_iff_ne_zero.mpr + ((Algebra.IsSeparable.isSeparable K x).aeval_derivative_ne_zero + (minpoly.aeval K x))).map f₂).mul_right_eq_zero, sub_eq_zero] at this + +variable [FormallyUnramified K A] [EssFiniteType K A] +variable [FormallyUnramified K L] [EssFiniteType K L] + +theorem bijective_of_isAlgClosed_of_localRing + [IsAlgClosed K] [LocalRing A] : + Function.Bijective (algebraMap K A) := by + have := finite_of_free (R := K) (S := A) + have : IsArtinianRing A := isArtinian_of_tower K inferInstance + have hA : IsNilpotent (LocalRing.maximalIdeal A) := by + rw [← LocalRing.jacobson_eq_maximalIdeal ⊥] + · exact IsArtinianRing.isNilpotent_jacobson_bot + · exact bot_ne_top + have : Function.Bijective (Algebra.ofId K (A ⧸ LocalRing.maximalIdeal A)) := + ⟨RingHom.injective _, IsAlgClosed.algebraMap_surjective_of_isIntegral⟩ + let e : K ≃ₐ[K] A ⧸ LocalRing.maximalIdeal A := { + __ := Algebra.ofId K (A ⧸ LocalRing.maximalIdeal A) + __ := Equiv.ofBijective _ this } + let e' : A ⊗[K] (A ⧸ LocalRing.maximalIdeal A) ≃ₐ[A] A := + (Algebra.TensorProduct.congr AlgEquiv.refl e.symm).trans (Algebra.TensorProduct.rid K A A) + let f : A ⧸ LocalRing.maximalIdeal A →ₗ[A] A := e'.toLinearMap.comp (sec K A _) + have hf : (Algebra.ofId _ _).toLinearMap ∘ₗ f = LinearMap.id := by + dsimp [f] + rw [← LinearMap.comp_assoc, ← comp_sec K A] + congr 1 + apply LinearMap.restrictScalars_injective K + apply _root_.TensorProduct.ext' + intros r s + obtain ⟨s, rfl⟩ := e.surjective s + suffices s • (Ideal.Quotient.mk (LocalRing.maximalIdeal A)) r = r • e s by + simpa [ofId, e'] + simp [Algebra.smul_def, e, ofId, mul_comm] + have hf₁ : f 1 • (1 : A ⧸ LocalRing.maximalIdeal A) = 1 := by + rw [← algebraMap_eq_smul_one] + exact LinearMap.congr_fun hf 1 + have hf₂ : 1 - f 1 ∈ LocalRing.maximalIdeal A := by + rw [← Ideal.Quotient.eq_zero_iff_mem, map_sub, map_one, ← Ideal.Quotient.algebraMap_eq, + algebraMap_eq_smul_one, hf₁, sub_self] + have hf₃ : IsIdempotentElem (1 - f 1) := by + apply IsIdempotentElem.one_sub + rw [IsIdempotentElem, ← smul_eq_mul, ← map_smul, hf₁] + have hf₄ : f 1 = 1 := by + obtain ⟨n, hn⟩ := hA + have : (1 - f 1) ^ n = 0 := by + rw [← Ideal.mem_bot, ← Ideal.zero_eq_bot, ← hn] + exact Ideal.pow_mem_pow hf₂ n + rw [eq_comm, ← sub_eq_zero, ← hf₃.pow_succ_eq n, pow_succ, this, zero_mul] + refine Equiv.bijective ⟨algebraMap K A, ⇑e.symm ∘ ⇑(algebraMap A _), fun x ↦ by simp, fun x ↦ ?_⟩ + have : ⇑(algebraMap K A) = ⇑f ∘ ⇑e := by + ext k + conv_rhs => rw [← mul_one k, ← smul_eq_mul, Function.comp_apply, map_smul, + LinearMap.map_smul_of_tower, map_one, hf₄, ← algebraMap_eq_smul_one] + rw [this] + simp only [Function.comp_apply, AlgEquiv.apply_symm_apply, algebraMap_eq_smul_one, + map_smul, hf₄, smul_eq_mul, mul_one] + +theorem isField_of_isAlgClosed_of_localRing + [IsAlgClosed K] [LocalRing A] : IsField A := by + rw [LocalRing.isField_iff_maximalIdeal_eq, eq_bot_iff] + intro x hx + obtain ⟨x, rfl⟩ := (bijective_of_isAlgClosed_of_localRing K A).surjective x + show _ = 0 + rw [← (algebraMap K A).map_zero] + by_contra hx' + exact hx ((isUnit_iff_ne_zero.mpr + (fun e ↦ hx' ((algebraMap K A).congr_arg e))).map (algebraMap K A)) + +include K in +theorem isReduced_of_field : + IsReduced A := by + constructor + intro x hx + let f := (Algebra.TensorProduct.includeRight (R := K) (A := AlgebraicClosure K) (B := A)) + have : Function.Injective f := by + have : ⇑f = (LinearMap.rTensor A (Algebra.ofId K (AlgebraicClosure K)).toLinearMap).comp + (Algebra.TensorProduct.lid K A).symm.toLinearMap := by + ext x; simp [f] + rw [this] + suffices Function.Injective + (LinearMap.rTensor A (Algebra.ofId K (AlgebraicClosure K)).toLinearMap) by + exact this.comp (Algebra.TensorProduct.lid K A).symm.injective + apply Module.Flat.rTensor_preserves_injective_linearMap + exact (algebraMap K _).injective + apply this + rw [map_zero] + apply eq_zero_of_localization + intro M hM + have hy := (hx.map f).map (algebraMap _ (Localization.AtPrime M)) + generalize algebraMap _ (Localization.AtPrime M) (f x) = y at * + have := EssFiniteType.of_isLocalization (Localization.AtPrime M) M.primeCompl + have := of_isLocalization (Rₘ := Localization.AtPrime M) M.primeCompl + have := EssFiniteType.comp (AlgebraicClosure K) (AlgebraicClosure K ⊗[K] A) + (Localization.AtPrime M) + have := comp (AlgebraicClosure K) (AlgebraicClosure K ⊗[K] A) + (Localization.AtPrime M) + letI := (isField_of_isAlgClosed_of_localRing (AlgebraicClosure K) + (A := Localization.AtPrime M)).toField + exact hy.eq_zero + +theorem range_eq_top_of_isPurelyInseparable + [IsPurelyInseparable K L] : (algebraMap K L).range = ⊤ := by + classical + have : Nontrivial (L ⊗[K] L) := by + rw [← not_subsingleton_iff_nontrivial, ← rank_zero_iff (R := K), rank_tensorProduct', + mul_eq_zero, or_self, rank_zero_iff, not_subsingleton_iff_nontrivial] + infer_instance + rw [← top_le_iff] + intro x _ + obtain ⟨n, hn⟩ := IsPurelyInseparable.pow_mem K (ringExpChar K) x + have : ExpChar (L ⊗[K] L) (ringExpChar K) := by + refine expChar_of_injective_ringHom (algebraMap K _).injective (ringExpChar K) + have : (1 ⊗ₜ x - x ⊗ₜ 1 : L ⊗[K] L) ^ (ringExpChar K) ^ n = 0 := by + rw [sub_pow_expChar_pow, TensorProduct.tmul_pow, one_pow, TensorProduct.tmul_pow, one_pow] + obtain ⟨r, hr⟩ := hn + rw [← hr, algebraMap_eq_smul_one, TensorProduct.smul_tmul, sub_self] + have H : (1 ⊗ₜ x : L ⊗[K] L) = x ⊗ₜ 1 := by + have inst : IsReduced (L ⊗[K] L) := isReduced_of_field L _ + exact sub_eq_zero.mp (IsNilpotent.eq_zero ⟨_, this⟩) + by_cases h' : LinearIndependent K ![1, x] + · have h := h'.coe_range + let S := h.extend (Set.subset_univ _) + let a : S := ⟨1, h.subset_extend _ (by simp)⟩; have ha : Basis.extend h a = 1 := by simp + let b : S := ⟨x, h.subset_extend _ (by simp)⟩; have hb : Basis.extend h b = x := by simp + by_cases e : a = b + · obtain rfl : 1 = x := congr_arg Subtype.val e + exact ⟨1, map_one _⟩ + have := DFunLike.congr_fun + (DFunLike.congr_arg ((Basis.extend h).tensorProduct (Basis.extend h)).repr H) (a, b) + simp only [Basis.tensorProduct_repr_tmul_apply, ← ha, ← hb, Basis.repr_self, smul_eq_mul, + Finsupp.single_apply, e, Ne.symm e, ↓reduceIte, mul_one, mul_zero, one_ne_zero] at this + · rw [LinearIndependent.pair_iff] at h' + simp only [not_forall, not_and, exists_prop] at h' + obtain ⟨a, b, e, hab⟩ := h' + have : IsUnit b := by + rw [isUnit_iff_ne_zero] + rintro rfl + rw [zero_smul, ← algebraMap_eq_smul_one, add_zero, + (injective_iff_map_eq_zero' _).mp (algebraMap K L).injective] at e + cases hab e rfl + use (-this.unit⁻¹ * a) + rw [map_mul, ← Algebra.smul_def, algebraMap_eq_smul_one, eq_neg_iff_add_eq_zero.mpr e, + smul_neg, neg_smul, neg_neg, smul_smul, this.val_inv_mul, one_smul] + +theorem isSeparable : Algebra.IsSeparable K L := by + have := finite_of_free (R := K) (S := L) + rw [← separableClosure.eq_top_iff] + have := of_comp K (separableClosure K L) L + have := EssFiniteType.of_comp K (separableClosure K L) L + have := separableClosure.isPurelyInseparable K L + ext + show _ ↔ _ ∈ (⊤ : Subring _) + rw [← range_eq_top_of_isPurelyInseparable (separableClosure K L) L] + simp + +theorem iff_isSeparable (L) [Field L] [Algebra K L] [EssFiniteType K L] : + FormallyUnramified K L ↔ Algebra.IsSeparable K L := + ⟨fun _ ↦ isSeparable K L, fun _ ↦ of_isSeparable K L⟩ + +end Algebra.FormallyUnramified From 24a4b4a6b273a4248ae80cc6615ba363c807bfc3 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 3 Oct 2024 17:44:04 +0000 Subject: [PATCH 158/174] =?UTF-8?q?refactor(AlgebraicGeometry):=20Introduc?= =?UTF-8?q?e=20`Scheme.toSpec=CE=93`=20(#15082)?= 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> Co-authored-by: Yury G. Kudryashov --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 25 +++---- .../GammaSpecAdjunction.lean | 67 ++++++++++--------- .../AlgebraicGeometry/Morphisms/Affine.lean | 3 +- .../Morphisms/QuasiSeparated.lean | 6 +- Mathlib/AlgebraicGeometry/Pullbacks.lean | 4 +- 5 files changed, 54 insertions(+), 51 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 56989b5600798..4b05220332013 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -52,24 +52,26 @@ deriving Category /-- A Scheme is affine if the canonical map `X ⟶ Spec Γ(X)` is an isomorphism. -/ class IsAffine (X : Scheme) : Prop where - affine : IsIso (ΓSpec.adjunction.unit.app X) + affine : IsIso X.toSpecΓ attribute [instance] IsAffine.affine +instance (X : Scheme.{u}) [IsAffine X] : IsIso (ΓSpec.adjunction.unit.app X) := @IsAffine.affine X _ + /-- The canonical isomorphism `X ≅ Spec Γ(X)` for an affine scheme. -/ @[simps! (config := .lemmasOnly) hom] def Scheme.isoSpec (X : Scheme) [IsAffine X] : X ≅ Spec Γ(X, ⊤) := - asIso (ΓSpec.adjunction.unit.app X) + asIso X.toSpecΓ @[reassoc] theorem Scheme.isoSpec_hom_naturality {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : X.isoSpec.hom ≫ Spec.map (f.app ⊤) = f ≫ Y.isoSpec.hom := by - simp only [isoSpec, asIso_hom, ΓSpec.adjunction_unit_naturality] + simp only [isoSpec, asIso_hom, Scheme.toSpecΓ_naturality] @[reassoc] theorem Scheme.isoSpec_inv_naturality {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : Spec.map (f.app ⊤) ≫ Y.isoSpec.inv = X.isoSpec.inv ≫ f := by - rw [Iso.eq_inv_comp, isoSpec, asIso_hom, ← ΓSpec.adjunction_unit_naturality_assoc, isoSpec, + rw [Iso.eq_inv_comp, isoSpec, asIso_hom, ← Scheme.toSpecΓ_naturality_assoc, isoSpec, asIso_inv, IsIso.hom_inv_id, Category.comp_id] /-- Construct an affine scheme from a scheme and the information that it is affine. @@ -228,7 +230,7 @@ theorem iSup_affineOpens_eq_top (X : Scheme) : ⨆ i : X.affineOpens, (i : X.Ope theorem Scheme.map_PrimeSpectrum_basicOpen_of_affine (X : Scheme) [IsAffine X] (f : Scheme.Γ.obj (op X)) : X.isoSpec.hom ⁻¹ᵁ PrimeSpectrum.basicOpen f = X.basicOpen f := - ΓSpec.adjunction_unit_map_basicOpen _ _ + Scheme.toSpecΓ_preimage_basicOpen _ _ theorem isBasis_basicOpen (X : Scheme) [IsAffine X] : Opens.IsBasis (Set.range (X.basicOpen : Γ(X, ⊤) → X.Opens)) := by @@ -240,10 +242,10 @@ theorem isBasis_basicOpen (X : Scheme) [IsAffine X] : constructor · rintro ⟨_, ⟨x, rfl⟩, rfl⟩ refine ⟨_, ⟨_, ⟨x, rfl⟩, rfl⟩, ?_⟩ - exact congr_arg Opens.carrier (ΓSpec.adjunction_unit_map_basicOpen _ _) + exact congr_arg Opens.carrier (Scheme.toSpecΓ_preimage_basicOpen _ _) · rintro ⟨_, ⟨_, ⟨x, rfl⟩, rfl⟩, rfl⟩ refine ⟨_, ⟨x, rfl⟩, ?_⟩ - exact congr_arg Opens.carrier (ΓSpec.adjunction_unit_map_basicOpen _ _).symm + exact congr_arg Opens.carrier (Scheme.toSpecΓ_preimage_basicOpen _ _).symm namespace IsAffineOpen @@ -356,10 +358,9 @@ theorem SpecΓIdentity_hom_app_fromSpec : (Scheme.ΓSpecIso Γ(X, U)).hom ≫ hU.fromSpec.app U = (Spec Γ(X, U)).presheaf.map (eqToHom hU.fromSpec_preimage_self).op := by simp only [fromSpec, Scheme.isoSpec, asIso_inv, Scheme.comp_coeBase, Opens.map_comp_obj, - ΓSpecIso_obj_hom, Scheme.Opens.topIso_inv, Opens.map_top, Functor.id_obj, Functor.comp_obj, - Functor.rightOp_obj, Scheme.Γ_obj, unop_op, Scheme.Spec_obj, Scheme.Opens.topIso_hom, - Scheme.comp_app, Scheme.Opens.ι_app_self, Category.assoc, ← Functor.map_comp_assoc, ← op_comp, - eqToHom_trans, Scheme.Opens.eq_presheaf_map_eqToHom, Scheme.Hom.naturality_assoc, + ΓSpecIso_obj_hom, Scheme.Opens.topIso_inv, Opens.map_top, Scheme.Opens.topIso_hom, + Scheme.comp_app, Scheme.Opens.ι_app_self, unop_op, Category.assoc, ← Functor.map_comp_assoc, ← + op_comp, eqToHom_trans, Scheme.Opens.eq_presheaf_map_eqToHom, Scheme.Hom.naturality_assoc, Scheme.inv_app_top, IsIso.hom_inv_id_assoc] simp only [eqToHom_op, eqToHom_map, Spec.map_eqToHom, eqToHom_unop, Scheme.Spec_map_presheaf_map_eqToHom, eqToHom_trans] @@ -683,7 +684,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, ⊤)) : - (ΓSpec.adjunction.unit.app X).val.base ⁻¹' PrimeSpectrum.zeroLocus s = X.zeroLocus s := + X.toSpecΓ.val.base ⁻¹' PrimeSpectrum.zeroLocus s = X.zeroLocus s := LocallyRingedSpace.toΓSpec_preimage_zeroLocus_eq s open ConcreteCategory diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 927e5943d373d..6e19f5f554874 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -403,14 +403,13 @@ theorem adjunction_counit_app' {R : CommRingCatᵒᵖ} : theorem adjunction_counit_app {R : CommRingCatᵒᵖ} : ΓSpec.adjunction.counit.app R = (Scheme.ΓSpecIso (unop R)).inv.op := rfl --- This is not a simp lemma to respect the abstraction -theorem adjunction_unit_app {X : Scheme} : - ΓSpec.adjunction.unit.app X = locallyRingedSpaceAdjunction.unit.app X.1 := rfl +/-- The canonical map `X ⟶ Spec Γ(X, ⊤)`. This is the unit of the `Γ-Spec` adjunction. -/ +def _root_.AlgebraicGeometry.Scheme.toSpecΓ (X : Scheme.{u}) : X ⟶ Spec Γ(X, ⊤) := + ΓSpec.adjunction.unit.app X -@[reassoc (attr := simp)] -theorem adjunction_unit_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : - f ≫ ΓSpec.adjunction.unit.app Y = ΓSpec.adjunction.unit.app X ≫ Spec.map (f.app ⊤) := - ΓSpec.adjunction.unit.naturality f +@[simp] +theorem adjunction_unit_app {X : Scheme} : + ΓSpec.adjunction.unit.app X = X.toSpecΓ := rfl instance isIso_locallyRingedSpaceAdjunction_counit : IsIso.{u + 1, u + 1} locallyRingedSpaceAdjunction.counit := @@ -422,56 +421,62 @@ instance isIso_adjunction_counit : IsIso ΓSpec.adjunction.counit := by rw [adjunction_counit_app] infer_instance +end ΓSpec + +@[reassoc (attr := simp)] +theorem Scheme.toSpecΓ_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : + f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.app ⊤) := + ΓSpec.adjunction.unit.naturality f + @[simp] -theorem adjunction_unit_app_app_top (X : Scheme.{u}) : - (ΓSpec.adjunction.unit.app X).app ⊤ = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by +theorem Scheme.toSpecΓ_app_top (X : Scheme.{u}) : + X.toSpecΓ.app ⊤ = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by have := ΓSpec.adjunction.left_triangle_components X dsimp at this rw [← IsIso.eq_comp_inv] at this - simp only [adjunction_counit_app, Functor.id_obj, Functor.comp_obj, Functor.rightOp_obj, + simp only [ΓSpec.adjunction_counit_app, Functor.id_obj, Functor.comp_obj, Functor.rightOp_obj, Scheme.Γ_obj, Category.id_comp] at this rw [← Quiver.Hom.op_inj.eq_iff, this, ← op_inv, IsIso.Iso.inv_inv] @[simp] theorem SpecMap_ΓSpecIso_hom (R : CommRingCat.{u}) : - Spec.map ((Scheme.ΓSpecIso R).hom) = adjunction.unit.app (Spec R) := by + Spec.map ((Scheme.ΓSpecIso R).hom) = (Spec R).toSpecΓ := by have := ΓSpec.adjunction.right_triangle_components (op R) dsimp at this rwa [← IsIso.eq_comp_inv, Category.id_comp, ← Spec.map_inv, IsIso.Iso.inv_inv, eq_comm] at this -lemma adjunction_unit_map_basicOpen (X : Scheme.{u}) (r : Γ(X, ⊤)) : - (ΓSpec.adjunction.unit.app X ⁻¹ᵁ (PrimeSpectrum.basicOpen r)) = X.basicOpen r := by - rw [← basicOpen_eq_of_affine] - erw [Scheme.preimage_basicOpen] +lemma Scheme.toSpecΓ_preimage_basicOpen (X : Scheme.{u}) (r : Γ(X, ⊤)) : + X.toSpecΓ ⁻¹ᵁ (PrimeSpectrum.basicOpen r) = X.basicOpen r := by + rw [← basicOpen_eq_of_affine, Scheme.preimage_basicOpen] congr - rw [ΓSpec.adjunction_unit_app_app_top] + rw [Scheme.toSpecΓ_app_top] exact Iso.inv_hom_id_apply _ _ -theorem toOpen_unit_app_val_c_app {X : Scheme.{u}} (U) : - StructureSheaf.toOpen _ _ ≫ (ΓSpec.adjunction.unit.app X).val.c.app U = +-- Warning: this LHS of this lemma breaks the structure-sheaf abstraction. +@[reassoc (attr := simp)] +theorem toOpen_toSpecΓ_app {X : Scheme.{u}} (U) : + StructureSheaf.toOpen _ _ ≫ X.toSpecΓ.app U = X.presheaf.map (homOfLE (by exact le_top)).op := by rw [← StructureSheaf.toOpen_res _ _ _ (homOfLE le_top), Category.assoc, - NatTrans.naturality _ (homOfLE (le_top (a := U.unop))).op] + NatTrans.naturality _ (homOfLE (le_top (a := U))).op] show (ΓSpec.adjunction.counit.app (Scheme.Γ.rightOp.obj X)).unop ≫ (Scheme.Γ.rightOp.map (ΓSpec.adjunction.unit.app X)).unop ≫ _ = _ rw [← Category.assoc, ← unop_comp, ΓSpec.adjunction.left_triangle_components] dsimp exact Category.id_comp _ --- Warning: this LHS of this lemma breaks the structure-sheaf abstraction. -@[reassoc (attr := simp)] -theorem toOpen_unit_app_val_c_app' {X : Scheme.{u}} (U : Opens (PrimeSpectrum Γ(X, ⊤))) : - toOpen Γ(X, ⊤) U ≫ (adjunction.unit.app X).app U = - X.presheaf.map (homOfLE (by exact le_top)).op := - ΓSpec.toOpen_unit_app_val_c_app (op U) - -end ΓSpec - theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : X.Opens) : (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).app ⊤ ≫ - (ΓSpec.adjunction.unit.app U).app ⊤ ≫ U.topIso.hom := by - rw [ΓSpec.adjunction_unit_app_app_top] -- why can't simp find this - simp + U.toScheme.toSpecΓ.app ⊤ ≫ U.topIso.hom := by simp + +@[deprecated (since := "2024-07-24")] +alias ΓSpec.adjunction_unit_naturality := Scheme.toSpecΓ_naturality +@[deprecated (since := "2024-07-24")] +alias ΓSpec.adjunction_unit_naturality_assoc := Scheme.toSpecΓ_naturality_assoc +@[deprecated (since := "2024-07-24")] +alias ΓSpec.adjunction_unit_app_app_top := Scheme.toSpecΓ_app_top +@[deprecated (since := "2024-07-24")] +alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen /-! Immediate consequences of the adjunction. -/ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index 9cf21c35955ab..c990b0ee818d7 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -122,8 +122,7 @@ lemma isAffine_of_isAffineOpen_basicOpen (s : Set Γ(X, ⊤)) simp only [← basicOpen_eq_of_affine] exact (isAffineOpen_top (Scheme.Spec.obj (op _))).basicOpen _ · rw [PrimeSpectrum.iSup_basicOpen_eq_top_iff, Subtype.range_coe_subtype, Set.setOf_mem_eq, hs] - · show IsAffineOpen (ΓSpec.adjunction.unit.app X ⁻¹ᵁ PrimeSpectrum.basicOpen i.1) - rw [ΓSpec.adjunction_unit_map_basicOpen] + · rw [Scheme.toSpecΓ_preimage_basicOpen] exact hs₂ _ i.2 · simp only [Functor.comp_obj, Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, id_eq, eq_mpr_eq_cast, Functor.id_obj, Opens.map_top, morphismRestrict_app] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index 83373737cbce2..fc114dd60f6a8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -353,7 +353,7 @@ 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 refine @IsIso.of_isIso_comp_right _ _ _ _ _ _ (X.presheaf.map - (eqToHom (ΓSpec.adjunction_unit_map_basicOpen _ _).symm).op) _ ?_ + (eqToHom (Scheme.toSpecΓ_preimage_basicOpen _ _).symm).op) _ ?_ rw [ConcreteCategory.isIso_iff_bijective, CommRingCat.forget_map] apply (config := { allowSynthFailures := true }) IsLocalization.bijective · exact StructureSheaf.IsLocalization.to_basicOpen _ _ @@ -361,8 +361,6 @@ theorem isIso_ΓSpec_adjunction_unit_app_basicOpen {X : Scheme} [CompactSpace X] · exact isCompact_univ · exact isQuasiSeparated_univ · rw [← CommRingCat.comp_eq_ring_hom_comp] - simp [RingHom.algebraMap_toAlgebra] - rw [ΓSpec.toOpen_unit_app_val_c_app'_assoc, ← Functor.map_comp] - rfl + simp [RingHom.algebraMap_toAlgebra, ← Functor.map_comp] end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 36ae5c042ab8c..2dcfa434d2017 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -453,8 +453,8 @@ instance isAffine_of_isAffine_isAffine_isAffine {X Y Z : Scheme} IsAffine (pullback f g) := isAffine_of_isIso (pullback.map f g (Spec.map (Γ.map f.op)) (Spec.map (Γ.map g.op)) - (ΓSpec.adjunction.unit.app X) (ΓSpec.adjunction.unit.app Y) (ΓSpec.adjunction.unit.app Z) - (ΓSpec.adjunction.unit.naturality f) (ΓSpec.adjunction.unit.naturality g) ≫ + X.toSpecΓ Y.toSpecΓ Z.toSpecΓ + (Scheme.toSpecΓ_naturality f) (Scheme.toSpecΓ_naturality g) ≫ (PreservesPullback.iso Scheme.Spec _ _).inv) /-- Given an open cover `{ Xᵢ }` of `X`, then `X ×[Z] Y` is covered by `Xᵢ ×[Z] Y`. -/ From 2d5610a78a9c94b0a15ec25e14a41a88bac202d0 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 3 Oct 2024 17:44:05 +0000 Subject: [PATCH 159/174] feat(RingTheory/Unramified): Formally unramified product of rings (#15141) --- Mathlib.lean | 1 + Mathlib/Algebra/Ring/Idempotents.lean | 9 +++ Mathlib/RingTheory/Unramified/Basic.lean | 46 +++++++---- Mathlib/RingTheory/Unramified/Pi.lean | 98 ++++++++++++++++++++++++ 4 files changed, 139 insertions(+), 15 deletions(-) create mode 100644 Mathlib/RingTheory/Unramified/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index b1d7ca948d1d3..58afee637aa81 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4091,6 +4091,7 @@ 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 import Mathlib.RingTheory.Valuation.AlgebraInstances import Mathlib.RingTheory.Valuation.Basic import Mathlib.RingTheory.Valuation.ExtendToLocalization diff --git a/Mathlib/Algebra/Ring/Idempotents.lean b/Mathlib/Algebra/Ring/Idempotents.lean index 415b47517507d..80a2c2e0be568 100644 --- a/Mathlib/Algebra/Ring/Idempotents.lean +++ b/Mathlib/Algebra/Ring/Idempotents.lean @@ -5,6 +5,7 @@ Authors: Christopher Hoskin -/ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Commute.Defs +import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Ring.Defs import Mathlib.Data.Subtype import Mathlib.Order.Notation @@ -49,6 +50,10 @@ theorem mul_of_commute {p q : S} (h : Commute p q) (h₁ : IsIdempotentElem p) (h₂ : IsIdempotentElem q) : IsIdempotentElem (p * q) := by rw [IsIdempotentElem, mul_assoc, ← mul_assoc q, ← h.eq, mul_assoc p, h₂.eq, ← mul_assoc, h₁.eq] +lemma mul {M} [CommSemigroup M] {e₁ e₂ : M} + (he₁ : IsIdempotentElem e₁) (he₂ : IsIdempotentElem e₂) : IsIdempotentElem (e₁ * e₂) := + he₁.mul_of_commute (.all e₁ e₂) he₂ + theorem zero : IsIdempotentElem (0 : M₀) := mul_zero _ @@ -83,6 +88,10 @@ theorem iff_eq_zero_or_one {p : G₀} : IsIdempotentElem p ↔ p = 0 ∨ p = 1 : h.elim (fun hp => hp.symm ▸ zero) fun hp => hp.symm ▸ one exact mul_left_cancel₀ hp (h.trans (mul_one p).symm) +lemma map {M N F} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] {e : M} + (he : IsIdempotentElem e) (f : F) : IsIdempotentElem (f e) := by + rw [IsIdempotentElem, ← map_mul, he.eq] + /-! ### Instances on `Subtype IsIdempotentElem` -/ diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index af45772e8e3eb..baf22e56e7160 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -104,23 +104,12 @@ theorem lift_unique' [FormallyUnramified R A] {C : Type u} [CommRing C] (g₁ g₂ : A →ₐ[R] B) (h : f.comp g₁ = f.comp g₂) : g₁ = g₂ := FormallyUnramified.ext' _ hf g₁ g₂ (AlgHom.congr_fun h) -end - -section OfEquiv - -variable {R : Type u} [CommSemiring R] -variable {A B : Type u} [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] - -theorem of_equiv [FormallyUnramified R A] (e : A ≃ₐ[R] B) : - FormallyUnramified R B := by +instance : FormallyUnramified R R := by constructor - 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] + intros B _ _ _ _ f₁ f₂ _ + exact Subsingleton.elim _ _ -end OfEquiv +end section Comp @@ -155,6 +144,33 @@ theorem of_comp [FormallyUnramified R B] : FormallyUnramified A B := by 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] + +/-- 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 + intro Q _ _ I hI f₁ f₂ e + ext x + obtain ⟨x, rfl⟩ := H x + rw [← AlgHom.comp_apply, ← AlgHom.comp_apply] + congr 1 + apply FormallyUnramified.comp_injective I hI + ext x; exact DFunLike.congr_fun e (f x) + +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 + +end of_surjective + section BaseChange open scoped TensorProduct diff --git a/Mathlib/RingTheory/Unramified/Pi.lean b/Mathlib/RingTheory/Unramified/Pi.lean new file mode 100644 index 0000000000000..dc1d2ec8c816b --- /dev/null +++ b/Mathlib/RingTheory/Unramified/Pi.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Unramified.Basic + +/-! + +# Formal-unramification of finite products of rings + +## Main result + +- `Algebra.FormallyUnramified.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. + +-/ + +namespace Algebra.FormallyUnramified + +universe u v + +variable {R : Type max u v} {I : Type v} [Finite I] (f : I → Type max u v) +variable [CommRing R] [∀ i, CommRing (f i)] [∀ i, Algebra R (f i)] + +theorem pi_iff : + FormallyUnramified R (∀ i, f i) ↔ ∀ i, FormallyUnramified R (f i) := by + classical + cases nonempty_fintype I + constructor + · intro _ i + exact FormallyUnramified.of_surjective (Pi.evalAlgHom R f i) (Function.surjective_eval i) + · intro H + constructor + intros B _ _ J hJ f₁ f₂ e + ext g + rw [← Finset.univ_sum_single g, map_sum, map_sum] + refine Finset.sum_congr rfl ?_ + rintro x - + have hf : ∀ x, f₁ x - f₂ x ∈ J := by + intro g + rw [← Ideal.Quotient.eq_zero_iff_mem, map_sub, sub_eq_zero] + exact AlgHom.congr_fun e g + let e : ∀ i, f i := Pi.single x 1 + have he : IsIdempotentElem e := by simp [IsIdempotentElem, e, ← Pi.single_mul] + have h₁ : (f₁ e) * (1 - f₂ e) = 0 := by + rw [← Ideal.mem_bot, ← hJ, ← ((he.map f₁).mul (he.map f₂).one_sub).eq, ← pow_two] + apply Ideal.pow_mem_pow + convert Ideal.mul_mem_left _ (f₁ e) (hf e) using 1 + rw [mul_sub, mul_sub, mul_one, (he.map f₁).eq] + have h₂ : (f₂ e) * (1 - f₁ e) = 0 := by + rw [← Ideal.mem_bot, ← hJ, ← ((he.map f₂).mul (he.map f₁).one_sub).eq, ← pow_two] + apply Ideal.pow_mem_pow + convert Ideal.mul_mem_left _ (-f₂ e) (hf e) using 1 + rw [neg_mul, mul_sub, mul_sub, mul_one, neg_sub, (he.map f₂).eq] + have H : f₁ e = f₂ e := by + trans f₁ e * f₂ e + · rw [← sub_eq_zero, ← h₁, mul_sub, mul_one] + · rw [eq_comm, ← sub_eq_zero, ← h₂, mul_sub, mul_one, mul_comm] + let J' := Ideal.span {1 - f₁ e} + let f₁' : f x →ₐ[R] B ⧸ J' := by + apply AlgHom.ofLinearMap + (((Ideal.Quotient.mkₐ R J').comp f₁).toLinearMap.comp (LinearMap.single _ _ x)) + · simp only [AlgHom.comp_toLinearMap, LinearMap.coe_comp, LinearMap.coe_single, + Function.comp_apply, AlgHom.toLinearMap_apply, Ideal.Quotient.mkₐ_eq_mk] + rw [eq_comm, ← sub_eq_zero, ← (Ideal.Quotient.mk J').map_one, ← map_sub, + Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] + · intros r s; simp [Pi.single_mul] + let f₂' : f x →ₐ[R] B ⧸ J' := by + apply AlgHom.ofLinearMap + (((Ideal.Quotient.mkₐ R J').comp f₂).toLinearMap.comp (LinearMap.single _ _ x)) + · simp only [AlgHom.comp_toLinearMap, LinearMap.coe_comp, LinearMap.coe_single, + Function.comp_apply, AlgHom.toLinearMap_apply, Ideal.Quotient.mkₐ_eq_mk] + rw [eq_comm, ← sub_eq_zero, ← (Ideal.Quotient.mk J').map_one, ← map_sub, + Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton, H] + · intros r s; simp [Pi.single_mul] + suffices f₁' = f₂' by + have := AlgHom.congr_fun this (g x) + simp only [AlgHom.comp_toLinearMap, AlgHom.ofLinearMap_apply, LinearMap.coe_comp, + LinearMap.coe_single, Function.comp_apply, AlgHom.toLinearMap_apply, ← map_sub, + Ideal.Quotient.mkₐ_eq_mk, ← sub_eq_zero (b := Ideal.Quotient.mk J' _), sub_zero, f₁', f₂', + Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton, J'] at this + obtain ⟨c, hc⟩ := this + apply_fun (f₁ e * ·) at hc + rwa [← mul_assoc, mul_sub, mul_sub, mul_one, (he.map f₁).eq, sub_self, zero_mul, + ← map_mul, H, ← map_mul, ← Pi.single_mul, one_mul, sub_eq_zero] at hc + apply FormallyUnramified.comp_injective (I := J.map (algebraMap _ _)) + · rw [← Ideal.map_pow, hJ, Ideal.map_bot] + · ext r + rw [← sub_eq_zero] + simp only [Ideal.Quotient.algebraMap_eq, AlgHom.coe_comp, Ideal.Quotient.mkₐ_eq_mk, + Function.comp_apply, ← map_sub, Ideal.Quotient.eq_zero_iff_mem, f₁', f₂', + AlgHom.comp_toLinearMap, AlgHom.ofLinearMap_apply, LinearMap.coe_comp, + LinearMap.coe_single, Function.comp_apply, AlgHom.toLinearMap_apply, + Ideal.Quotient.mkₐ_eq_mk] + exact Ideal.mem_map_of_mem (Ideal.Quotient.mk J') (hf (Pi.single x r)) + +end Algebra.FormallyUnramified From 78d98550c5a604e7c2cde8aa3378f1f3a9eae733 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Thu, 3 Oct 2024 17:44:06 +0000 Subject: [PATCH 160/174] feat(CategoryTheory/Limits): Some infrastructure for the preservation of finite (co)limits (#17156) --- Mathlib/CategoryTheory/Limits/Preserves/Finite.lean | 11 ++++++++++- .../Limits/Preserves/FunctorCategory.lean | 13 +++++++++++++ 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean index 87c35359a609f..b368aca7e3616 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean @@ -84,6 +84,11 @@ def compPreservesFiniteLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits [PreservesFiniteLimits G] : PreservesFiniteLimits (F ⋙ G) := ⟨fun _ _ _ => inferInstance⟩ +/-- Transfer preservation of finite limits along a natural isomorphism in the functor. -/ +def preservesFiniteLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteLimits F] : + PreservesFiniteLimits G where + preservesFiniteLimits _ _ _ := preservesLimitsOfShapeOfNatIso h + /- Porting note: adding this class because quantified classes don't behave well [#2764](https://github.com/leanprover-community/mathlib4/pull/2764) -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` @@ -225,6 +230,11 @@ def compPreservesFiniteColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColi [PreservesFiniteColimits G] : PreservesFiniteColimits (F ⋙ G) := ⟨fun _ _ _ => inferInstance⟩ +/-- Transfer preservation of finite colimits along a natural isomorphism in the functor. -/ +def preservesFiniteColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteColimits F] : + PreservesFiniteColimits G where + preservesFiniteColimits _ _ _ := preservesColimitsOfShapeOfNatIso h + /- Porting note: adding this class because quantified classes don't behave well [#2764](https://github.com/leanprover-community/mathlib4/pull/2764) -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` @@ -248,7 +258,6 @@ noncomputable instance compPreservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) noncomputable instance (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteCoproducts F where preserves _ _ := inferInstance - /-- A functor is said to reflect finite colimits, if it reflects all colimits of shape `J`, where `J : Type` is a finite category. diff --git a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean index 4b4cde050f773..0d33dbbf72e5a 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts +import Mathlib.CategoryTheory.Limits.Preserves.Finite import Mathlib.CategoryTheory.Limits.Yoneda import Mathlib.CategoryTheory.Limits.Presheaf @@ -104,4 +105,16 @@ noncomputable def preservesLimitOfLanPreservesLimit {C D : Type u} [SmallCategor apply @preservesLimitsOfShapeOfReflectsOfPreserves _ _ _ _ _ _ _ _ F yoneda ?_ exact preservesLimitsOfShapeOfNatIso (Presheaf.compYonedaIsoYonedaCompLan F).symm +/-- `F : C ⥤ D ⥤ E` preserves finite limits if it does for each `d : D`. -/ +def preservesFiniteLimitsOfEvaluation {D : Type*} [Category D] {E : Type*} [Category E] + (F : C ⥤ D ⥤ E) (h : ∀ d : D, PreservesFiniteLimits (F ⋙ (evaluation D E).obj d)) : + PreservesFiniteLimits F := + ⟨fun J _ _ => preservesLimitsOfShapeOfEvaluation F J fun k => (h k).preservesFiniteLimits _⟩ + +/-- `F : C ⥤ D ⥤ E` preserves finite limits if it does for each `d : D`. -/ +def preservesFiniteColimitsOfEvaluation {D : Type*} [Category D] {E : Type*} [Category E] + (F : C ⥤ D ⥤ E) (h : ∀ d : D, PreservesFiniteColimits (F ⋙ (evaluation D E).obj d)) : + PreservesFiniteColimits F := + ⟨fun J _ _ => preservesColimitsOfShapeOfEvaluation F J fun k => (h k).preservesFiniteColimits _⟩ + end CategoryTheory From 0c13e18343c06775d5f2121a433f70392b7e2caa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 19:18:22 +0000 Subject: [PATCH 161/174] chore: Rename `FiniteDimensional.finrank` to `Module.finrank` (#17192) The namespacing in the linear algebra/affine spaces part of the library is a bit suboptimal. Here are a few arguments for why `FiniteDimensional.finrank` should be renamed to `Module.finrank`: 1. Other objects in maths can be finite-dimensional and have a rank. 2. `finrank` is directly analogous to `Module.rank`, they have lemmas in common (eg `finrank_eq_rank`, which is currently in neither of the declarations' namespaces) and have common series of lemmas that relate to other declarations in the `Module` namespace, eg `Module.Finite`. 3. We are in the process of replacing `FiniteDimensional` by `Module.Finite`, so the `FiniteDimensional` namespace will soon be obsolete. 4. It is a shorter and clearer name. Also make `Module.Finite` protected to avoid clashes with `Finite` --- Archive/Imo/Imo2019Q2.lean | 2 +- Archive/Sensitivity.lean | 2 +- .../Algebra/Subalgebra/IsSimpleOrder.lean | 2 +- Mathlib/Algebra/Algebra/Subalgebra/Rank.lean | 2 +- Mathlib/Algebra/Category/ModuleCat/Free.lean | 10 +- .../Algebra/Category/ModuleCat/Simple.lean | 2 +- Mathlib/Algebra/Lie/CartanExists.lean | 6 +- Mathlib/Algebra/Lie/InvariantForm.lean | 6 +- Mathlib/Algebra/Lie/Rank.lean | 12 +-- Mathlib/Algebra/Lie/TraceForm.lean | 4 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 4 +- Mathlib/Algebra/Lie/Weights/Chain.lean | 2 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 2 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 2 +- .../Algebra/Module/LinearMap/Polynomial.lean | 8 +- Mathlib/Algebra/Module/Submodule/Map.lean | 2 +- Mathlib/Algebra/Module/ZLattice/Basic.lean | 4 +- Mathlib/Algebra/Module/ZLattice/Covolume.lean | 2 +- Mathlib/Algebra/Polynomial/Module/AEval.lean | 7 +- Mathlib/Algebra/Quaternion.lean | 6 +- .../EllipticCurve/Group.lean | 2 +- .../Calculus/BumpFunction/Convolution.lean | 2 +- .../BumpFunction/FiniteDimension.lean | 2 +- .../Calculus/BumpFunction/Normed.lean | 2 +- .../Calculus/ContDiff/FiniteDimension.lean | 2 +- .../LineDeriv/IntegrationByParts.lean | 2 +- Mathlib/Analysis/Calculus/Rademacher.lean | 2 +- Mathlib/Analysis/Convex/Measure.lean | 4 +- Mathlib/Analysis/Convex/Normed.lean | 2 +- Mathlib/Analysis/Convex/Radon.lean | 2 +- .../Analysis/Distribution/SchwartzSpace.lean | 4 +- Mathlib/Analysis/Fourier/Inversion.lean | 2 +- .../Fourier/RiemannLebesgueLemma.lean | 2 +- .../FunctionalSpaces/SobolevInequality.lean | 2 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 2 +- .../InnerProductSpace/EuclideanDist.lean | 2 +- .../InnerProductSpace/GramSchmidtOrtho.lean | 2 +- .../InnerProductSpace/Orientation.lean | 2 +- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 12 +-- .../InnerProductSpace/Projection.lean | 8 +- .../Analysis/InnerProductSpace/Spectrum.lean | 2 +- .../Analysis/InnerProductSpace/TwoDim.lean | 4 +- .../Normed/Module/FiniteDimension.lean | 8 +- .../NormedSpace/HahnBanach/Extension.lean | 4 +- .../Gaussian/FourierTransform.lean | 10 +- .../SpecialFunctions/JapaneseBracket.lean | 2 +- Mathlib/CategoryTheory/Preadditive/Schur.lean | 2 +- .../Combinatorics/SimpleGraph/LapMatrix.lean | 4 +- Mathlib/Data/Complex/FiniteDimensional.lean | 9 +- Mathlib/Data/Matrix/Rank.lean | 6 +- Mathlib/FieldTheory/Adjoin.lean | 17 ++-- Mathlib/FieldTheory/Cardinality.lean | 4 +- Mathlib/FieldTheory/Finite/Basic.lean | 4 +- Mathlib/FieldTheory/Finite/GaloisField.lean | 4 +- Mathlib/FieldTheory/Finite/Polynomial.lean | 6 +- Mathlib/FieldTheory/Fixed.lean | 2 +- Mathlib/FieldTheory/Galois/Basic.lean | 2 +- .../IntermediateField/Algebraic.lean | 9 +- Mathlib/FieldTheory/IsPerfectClosure.lean | 2 +- Mathlib/FieldTheory/KummerExtension.lean | 10 +- Mathlib/FieldTheory/Minpoly/Field.lean | 6 +- Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean | 2 +- .../FieldTheory/PolynomialGaloisGroup.lean | 6 +- Mathlib/FieldTheory/PrimitiveElement.lean | 6 +- Mathlib/FieldTheory/PurelyInseparable.lean | 4 +- Mathlib/FieldTheory/SeparableClosure.lean | 2 +- Mathlib/FieldTheory/SeparableDegree.lean | 14 +-- Mathlib/FieldTheory/Tower.lean | 2 +- .../Euclidean/Angle/Oriented/Affine.lean | 2 +- .../Euclidean/Angle/Oriented/Basic.lean | 2 +- .../Euclidean/Angle/Oriented/RightAngle.lean | 4 +- .../Euclidean/Angle/Oriented/Rotation.lean | 8 +- Mathlib/Geometry/Euclidean/Angle/Sphere.lean | 2 +- Mathlib/Geometry/Euclidean/Basic.lean | 4 +- Mathlib/Geometry/Euclidean/Circumcenter.lean | 2 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 13 +-- Mathlib/Geometry/Euclidean/Sphere/Basic.lean | 2 +- Mathlib/Geometry/Euclidean/Triangle.lean | 2 +- Mathlib/Geometry/Manifold/BumpFunction.lean | 2 +- .../Geometry/Manifold/Instances/Sphere.lean | 2 +- .../Geometry/Manifold/PartitionOfUnity.lean | 2 +- .../Geometry/Manifold/WhitneyEmbedding.lean | 2 +- .../AffineSpace/FiniteDimensional.lean | 16 +-- .../BilinearForm/Orthogonal.lean | 2 +- Mathlib/LinearAlgebra/Charpoly/Basic.lean | 2 +- Mathlib/LinearAlgebra/Coevaluation.lean | 2 +- Mathlib/LinearAlgebra/Contraction.lean | 2 +- Mathlib/LinearAlgebra/Determinant.lean | 16 +-- .../Dimension/Constructions.lean | 22 ++--- .../LinearAlgebra/Dimension/DivisionRing.lean | 4 +- Mathlib/LinearAlgebra/Dimension/Finite.lean | 24 ++--- Mathlib/LinearAlgebra/Dimension/Finrank.lean | 10 +- Mathlib/LinearAlgebra/Dimension/Free.lean | 37 +++---- .../Dimension/FreeAndStrongRankCondition.lean | 6 +- .../LinearAlgebra/Dimension/Localization.lean | 4 +- .../LinearAlgebra/Dimension/RankNullity.lean | 2 +- .../Dimension/StrongRankCondition.lean | 34 +++---- Mathlib/LinearAlgebra/Dual.lean | 25 ++--- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean | 2 +- .../Eigenspace/Triangularizable.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 2 +- Mathlib/LinearAlgebra/FiniteDimensional.lean | 6 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 99 +++++++++---------- .../FreeModule/Finite/Matrix.lean | 16 +-- .../FreeModule/IdealQuotient.lean | 6 +- Mathlib/LinearAlgebra/FreeModule/Norm.lean | 2 +- .../Matrix/GeneralLinearGroup/Card.lean | 6 +- Mathlib/LinearAlgebra/Orientation.lean | 2 +- .../LinearAlgebra/Projectivization/Basic.lean | 2 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 4 +- .../LinearAlgebra/QuadraticForm/Complex.lean | 2 +- .../QuadraticForm/IsometryEquiv.lean | 6 +- Mathlib/LinearAlgebra/QuadraticForm/Real.lean | 8 +- Mathlib/LinearAlgebra/Semisimple.lean | 8 +- .../TensorProduct/Subalgebra.lean | 6 +- Mathlib/LinearAlgebra/Trace.lean | 2 +- .../Constructions/HaarToSphere.lean | 8 +- .../Covering/BesicovitchVectorSpace.lean | 2 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 4 +- .../Group/GeometryOfNumbers.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 2 +- .../Measure/Haar/InnerProductSpace.lean | 2 +- .../Measure/Haar/NormedSpace.lean | 6 +- .../MeasureTheory/Measure/Haar/OfBasis.lean | 2 +- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 2 +- .../Measure/Lebesgue/EqHaar.lean | 2 +- .../Measure/Lebesgue/VolumeOfBalls.lean | 6 +- Mathlib/NumberTheory/Cyclotomic/Basic.lean | 2 +- .../NumberTheory/Cyclotomic/Embeddings.lean | 2 +- .../Cyclotomic/PrimitiveRoots.lean | 6 +- Mathlib/NumberTheory/FunctionField.lean | 4 +- Mathlib/NumberTheory/NumberField/Basic.lean | 2 +- .../NumberField/CanonicalEmbedding/Basic.lean | 6 +- .../CanonicalEmbedding/ConvexBody.lean | 4 +- .../CanonicalEmbedding/FundamentalCone.lean | 2 +- .../NumberTheory/NumberField/ClassNumber.lean | 2 +- .../NumberField/Discriminant.lean | 2 +- .../NumberTheory/NumberField/Embeddings.lean | 14 +-- .../NumberField/EquivReindex.lean | 2 +- .../NumberField/FractionalIdeal.lean | 2 +- Mathlib/NumberTheory/NumberField/House.lean | 4 +- Mathlib/NumberTheory/NumberField/Norm.lean | 2 +- .../NumberField/Units/DirichletTheorem.lean | 6 +- Mathlib/NumberTheory/RamificationInertia.lean | 4 +- Mathlib/RepresentationTheory/Character.lean | 4 +- Mathlib/RepresentationTheory/FDRep.lean | 2 +- .../DedekindDomain/IntegralClosure.lean | 4 +- .../DiscreteValuationRing/TFAE.lean | 2 +- Mathlib/RingTheory/Discriminant.lean | 4 +- Mathlib/RingTheory/FiniteType.lean | 4 +- Mathlib/RingTheory/Finiteness.lean | 41 ++++---- .../RingTheory/Flat/EquationalCriterion.lean | 14 +-- Mathlib/RingTheory/Ideal/Cotangent.lean | 2 +- Mathlib/RingTheory/LittleWedderburn.lean | 2 +- Mathlib/RingTheory/LocalRing/Module.lean | 8 +- Mathlib/RingTheory/Noetherian.lean | 6 +- Mathlib/RingTheory/Norm/Basic.lean | 4 +- Mathlib/RingTheory/Norm/Defs.lean | 2 +- .../Polynomial/Eisenstein/IsIntegral.lean | 2 +- Mathlib/RingTheory/PowerBasis.lean | 6 +- Mathlib/RingTheory/SimpleModule.lean | 2 +- Mathlib/RingTheory/Trace/Basic.lean | 4 +- Mathlib/RingTheory/Trace/Defs.lean | 2 +- Mathlib/RingTheory/Valuation/Minpoly.lean | 2 +- Mathlib/RingTheory/WittVector/Isocrystal.lean | 4 +- .../Algebra/Module/FiniteDimension.lean | 6 +- .../MetricSpace/HausdorffDimension.lean | 6 +- 168 files changed, 477 insertions(+), 500 deletions(-) diff --git a/Archive/Imo/Imo2019Q2.lean b/Archive/Imo/Imo2019Q2.lean index db1650a221391..86353bd5b6d91 100644 --- a/Archive/Imo/Imo2019Q2.lean +++ b/Archive/Imo/Imo2019Q2.lean @@ -57,7 +57,7 @@ rather than more literally with `affineSegment`. -/ -open Affine Affine.Simplex EuclideanGeometry FiniteDimensional +open Affine Affine.Simplex EuclideanGeometry Module open scoped Affine EuclideanGeometry Real diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index d28ffea139084..5cae8d062537c 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -41,7 +41,7 @@ noncomputable section local notation "√" => Real.sqrt -open Function Bool LinearMap Fintype FiniteDimensional Module.DualBases +open Function Bool LinearMap Fintype Module Module.DualBases /-! ### The hypercube diff --git a/Mathlib/Algebra/Algebra/Subalgebra/IsSimpleOrder.lean b/Mathlib/Algebra/Algebra/Subalgebra/IsSimpleOrder.lean index 468781008d3b8..2c9016caf3b6c 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/IsSimpleOrder.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/IsSimpleOrder.lean @@ -12,7 +12,7 @@ If `A` is a domain, and a finite-dimensional algebra over a field `F`, with prim then there are no non-trivial `F`-subalgebras. -/ -open FiniteDimensional Submodule +open Module Submodule theorem Subalgebra.isSimpleOrder_of_finrank_prime (F A) [Field F] [Ring A] [IsDomain A] [Algebra F A] (hp : (finrank F A).Prime) : IsSimpleOrder (Subalgebra F A) := diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean b/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean index 1dac252bbf591..e4b3c91e4d894 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean @@ -18,7 +18,7 @@ satisfies strong rank condition, we put them into a separate file. -/ -open FiniteDimensional +open Module namespace Subalgebra diff --git a/Mathlib/Algebra/Category/ModuleCat/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Free.lean index 0cf3303dd9713..9f822325b9425 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Free.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Free.lean @@ -26,7 +26,7 @@ linear algebra, module, free -/ -open CategoryTheory +open CategoryTheory Module namespace ModuleCat @@ -177,11 +177,11 @@ theorem free_shortExact_rank_add [Module.Free R S.X₁] [Module.Free R S.X₃] theorem free_shortExact_finrank_add {n p : ℕ} [Module.Free R S.X₁] [Module.Free R S.X₃] [Module.Finite R S.X₁] [Module.Finite R S.X₃] - (hN : FiniteDimensional.finrank R S.X₁ = n) - (hP : FiniteDimensional.finrank R S.X₃ = p) + (hN : Module.finrank R S.X₁ = n) + (hP : Module.finrank R S.X₃ = p) [StrongRankCondition R] : - FiniteDimensional.finrank R S.X₂ = n + p := by - apply FiniteDimensional.finrank_eq_of_rank_eq + finrank R S.X₂ = n + p := by + apply finrank_eq_of_rank_eq rw [free_shortExact_rank_add hS', ← hN, ← hP] simp only [Nat.cast_add, finrank_eq_rank] diff --git a/Mathlib/Algebra/Category/ModuleCat/Simple.lean b/Mathlib/Algebra/Category/ModuleCat/Simple.lean index dcccbac0230b4..a36bd8151ab84 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Simple.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Simple.lean @@ -34,7 +34,7 @@ instance simple_of_isSimpleModule [IsSimpleModule R M] : Simple (of R M) := instance isSimpleModule_of_simple (M : ModuleCat R) [Simple M] : IsSimpleModule R M := simple_iff_isSimpleModule.mp (Simple.of_iso (ofSelfIso M)) -open FiniteDimensional +open Module attribute [local instance] moduleOfAlgebraModule isScalarTower_of_algebra_moduleCat diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index ae5ba4fc34d14..5ee149734e3b5 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -39,7 +39,7 @@ variable [Module.Finite K L] variable [Module.Finite R L] [Module.Free R L] variable [Module.Finite R M] [Module.Free R M] -open FiniteDimensional LieSubalgebra Module.Free Polynomial +open Module LieSubalgebra Module.Free Polynomial variable (K) @@ -117,7 +117,7 @@ section Field variable {K L : Type*} [Field K] [LieRing L] [LieAlgebra K L] [Module.Finite K L] -open FiniteDimensional LieSubalgebra LieSubmodule Polynomial Cardinal LieModule engel_isBot_of_isMin +open Module LieSubalgebra LieSubmodule Polynomial Cardinal LieModule engel_isBot_of_isMin #adaptation_note /-- otherwise there is a spurious warning on `contrapose!` below. -/ set_option linter.unusedVariables false in @@ -360,7 +360,7 @@ lemma exists_isCartanSubalgebra_engel_of_finrank_le_card (h : finrank K L ≤ #K suffices finrank K (engel K x) ≤ finrank K (engel K y) by suffices engel K y = engel K x from this.ge apply LieSubalgebra.to_submodule_injective - exact eq_of_le_of_finrank_le hyx this + exact Submodule.eq_of_le_of_finrank_le hyx this rw [(isRegular_iff_finrank_engel_eq_rank K x).mp hx] apply rank_le_finrank_engel diff --git a/Mathlib/Algebra/Lie/InvariantForm.lean b/Mathlib/Algebra/Lie/InvariantForm.lean index dea6e6da8c49d..e95fdc8b0f1bf 100644 --- a/Mathlib/Algebra/Lie/InvariantForm.lean +++ b/Mathlib/Algebra/Lie/InvariantForm.lean @@ -124,14 +124,14 @@ variable (hΦ_inv : Φ.lieInvariant L) (hΦ_refl : Φ.IsRefl) variable (hL : ∀ I : LieIdeal K L, IsAtom I → ¬IsLieAbelian I) include hΦ_nondeg hΦ_refl hL -open FiniteDimensional Submodule in +open Module Submodule in lemma orthogonal_isCompl_coe_submodule (I : LieIdeal K L) (hI : IsAtom I) : IsCompl I.toSubmodule (orthogonal Φ hΦ_inv I).toSubmodule := by rw [orthogonal_toSubmodule, LinearMap.BilinForm.isCompl_orthogonal_iff_disjoint hΦ_refl, ← orthogonal_toSubmodule _ hΦ_inv, ← LieSubmodule.disjoint_iff_coe_toSubmodule] exact orthogonal_disjoint Φ hΦ_nondeg hΦ_inv hL I hI -open FiniteDimensional Submodule in +open Module Submodule in lemma orthogonal_isCompl (I : LieIdeal K L) (hI : IsAtom I) : IsCompl I (orthogonal Φ hΦ_inv I) := by rw [LieSubmodule.isCompl_iff_coe_toSubmodule] @@ -151,7 +151,7 @@ lemma restrict_orthogonal_nondegenerate (I : LieIdeal K L) (hI : IsAtom I) : LinearMap.BilinForm.orthogonal_orthogonal hΦ_nondeg hΦ_refl] exact (orthogonal_isCompl_coe_submodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI).symm -open FiniteDimensional Submodule in +open Module Submodule in lemma atomistic : ∀ I : LieIdeal K L, sSup {J : LieIdeal K L | IsAtom J ∧ J ≤ I} = I := by intro I apply le_antisymm diff --git a/Mathlib/Algebra/Lie/Rank.lean b/Mathlib/Algebra/Lie/Rank.lean index 3fabb5568a84c..c9c5877344913 100644 --- a/Mathlib/Algebra/Lie/Rank.lean +++ b/Mathlib/Algebra/Lie/Rank.lean @@ -65,13 +65,13 @@ lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : rank R L M = (polyCharpoly φ b).natTrailingDegree := by apply nilRank_eq_polyCharpoly_natTrailingDegree -open FiniteDimensional +open Module include bₘ in lemma rank_le_card [Nontrivial R] : rank R L M ≤ Fintype.card ιₘ := nilRank_le_card _ bₘ -open FiniteDimensional +open Module lemma rank_le_finrank [Nontrivial R] : rank R L M ≤ finrank R M := nilRank_le_finrank _ @@ -103,7 +103,7 @@ section IsDomain variable (L) variable [IsDomain R] -open Cardinal FiniteDimensional MvPolynomial in +open Cardinal Module MvPolynomial in lemma exists_isRegular_of_finrank_le_card (h : finrank R M ≤ #R) : ∃ x : L, IsRegular R M x := LinearMap.exists_isNilRegular_of_finrank_le_card _ h @@ -138,7 +138,7 @@ lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : rank R L = (polyCharpoly (ad R L).toLinearMap b).natTrailingDegree := by apply nilRank_eq_polyCharpoly_natTrailingDegree -open FiniteDimensional +open Module include b in lemma rank_le_card [Nontrivial R] : rank R L ≤ Fintype.card ι := @@ -175,7 +175,7 @@ section IsDomain variable (L) variable [IsDomain R] -open Cardinal FiniteDimensional MvPolynomial in +open Cardinal Module MvPolynomial in lemma exists_isRegular_of_finrank_le_card (h : finrank R L ≤ #R) : ∃ x : L, IsRegular R x := LinearMap.exists_isNilRegular_of_finrank_le_card _ h @@ -191,7 +191,7 @@ namespace LieAlgebra variable (K : Type*) {L : Type*} [Field K] [LieRing L] [LieAlgebra K L] [Module.Finite K L] -open FiniteDimensional LieSubalgebra +open Module LieSubalgebra lemma finrank_engel (x : L) : finrank K (engel K x) = (ad K L x).charpoly.natTrailingDegree := diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index b7b439cf4ddfd..3e6c88e8e9ed1 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -38,7 +38,7 @@ variable (R K L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] local notation "φ" => LieModule.toEnd R L M open LinearMap (trace) -open Set FiniteDimensional +open Set Module namespace LieModule @@ -392,7 +392,7 @@ lemma killingForm_eq : end LieIdeal -open LieModule FiniteDimensional +open LieModule Module open Submodule (span subset_span) namespace LieModule diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 810246b2c95a8..bebee81c0475c 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -745,7 +745,7 @@ lemma iSup_genWeightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : simp_rw [Module.End.maxGenEigenspace_def] exact IsTriangularizable.iSup_eq_top x -open LinearMap FiniteDimensional in +open LinearMap Module in @[simp] lemma trace_toEnd_genWeightSpace [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] (χ : L → R) (x : L) : @@ -759,7 +759,7 @@ lemma trace_toEnd_genWeightSpace [IsDomain R] [IsPrincipalIdealRing R] section field -open FiniteDimensional +open Module variable (K) variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [LieAlgebra.IsNilpotent K L] diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index bf2671f9c4c4b..428b6ff73884c 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -41,7 +41,7 @@ We provide basic definitions and results to support `α`-chain techniques in thi -/ -open FiniteDimensional Function Set +open Module Function Set variable {R L : Type*} [CommRing R] [LieRing L] [LieAlgebra R L] (M : Type*) [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index a941003c52c63..7b0522b1c70b9 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -85,7 +85,7 @@ end IsKilling section Field -open FiniteDimensional LieModule Set +open Module LieModule Set open Submodule (span subset_span) variable [FiniteDimensional K L] (H : LieSubalgebra K L) [H.IsCartanSubalgebra] diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index a4da88d28a011..9798876d617f9 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -114,7 +114,7 @@ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R section FiniteDimensional -open FiniteDimensional +open Module variable [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] [LieAlgebra.IsNilpotent R L] diff --git a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean index 6f0b08d7782c3..823b874234f59 100644 --- a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean +++ b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean @@ -351,7 +351,7 @@ lemma polyCharpolyAux_basisIndep {ιM' : Type*} [Fintype ιM'] [DecidableEq ιM' end aux -open FiniteDimensional Matrix +open Module Matrix variable [Module.Free R M] [Module.Finite R M] (b : Basis ι R L) @@ -479,11 +479,11 @@ lemma polyCharpoly_coeff_nilRank_ne_zero : rw [nilRank_eq_polyCharpoly_natTrailingDegree _ b] apply polyCharpoly_coeff_nilRankAux_ne_zero -open FiniteDimensional Module.Free +open Module Module.Free lemma nilRank_le_card {ι : Type*} [Fintype ι] (b : Basis ι R M) : nilRank φ ≤ Fintype.card ι := by apply Polynomial.natTrailingDegree_le_of_ne_zero - rw [← FiniteDimensional.finrank_eq_card_basis b, ← polyCharpoly_natDegree φ (chooseBasis R L), + rw [← Module.finrank_eq_card_basis b, ← polyCharpoly_natDegree φ (chooseBasis R L), Polynomial.coeff_natDegree, (polyCharpoly_monic _ _).leadingCoeff] apply one_ne_zero @@ -538,7 +538,7 @@ section IsDomain variable [IsDomain R] -open Cardinal FiniteDimensional MvPolynomial Module.Free in +open Cardinal Module MvPolynomial Module.Free in lemma exists_isNilRegular_of_finrank_le_card (h : finrank R M ≤ #R) : ∃ x : L, IsNilRegular φ x := by let b := chooseBasis R L diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 1d1245ea336c6..0e08a969e1c07 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -104,7 +104,7 @@ theorem map_mono {f : F} {p p' : Submodule R M} : p ≤ p' → map f p ≤ map f image_subset _ @[simp] -theorem map_zero : map (0 : M →ₛₗ[σ₁₂] M₂) p = ⊥ := +protected theorem map_zero : map (0 : M →ₛₗ[σ₁₂] M₂) p = ⊥ := have : ∃ x : M, x ∈ p := ⟨0, p.zero_mem⟩ ext <| by simp [this, eq_comm] diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index 0ae1e6609864d..a436852b2ab88 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -384,7 +384,7 @@ end ZSpan section ZLattice -open Submodule FiniteDimensional ZSpan +open Submodule Module ZSpan -- TODO: generalize this class to other rings than `ℤ` /-- `L : Submodule ℤ E` where `E` is a vector space over a normed field `K` is a `ℤ`-lattice if @@ -562,7 +562,7 @@ variable {ι : Type*} [hs : IsZLattice K L] (b : Basis ι ℤ L) /-- Any `ℤ`-basis of `L` is also a `K`-basis of `E`. -/ def Basis.ofZLatticeBasis : Basis ι K E := by - have : Finite ℤ L := ZLattice.module_finite K L + have : Module.Finite ℤ L := ZLattice.module_finite K L have : Free ℤ L := ZLattice.module_free K L let e := Basis.indexEquiv (Free.chooseBasis ℤ L) b have : Fintype ι := Fintype.ofEquiv _ e diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index bec47d37fcc0b..04e559d80a384 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -29,7 +29,7 @@ noncomputable section namespace ZLattice -open Submodule MeasureTheory FiniteDimensional MeasureTheory Module +open Submodule MeasureTheory Module MeasureTheory Module section General diff --git a/Mathlib/Algebra/Polynomial/Module/AEval.lean b/Mathlib/Algebra/Polynomial/Module/AEval.lean index daa261800b354..17b934cb8564e 100644 --- a/Mathlib/Algebra/Polynomial/Module/AEval.lean +++ b/Mathlib/Algebra/Polynomial/Module/AEval.lean @@ -46,7 +46,8 @@ instance instAddCommMonoid : AddCommMonoid <| AEval R M a := inferInstanceAs (Ad instance instModuleOrig : Module R <| AEval R M a := inferInstanceAs (Module R M) -instance instFiniteOrig [Finite R M] : Finite R <| AEval R M a := inferInstanceAs (Finite R M) +instance instFiniteOrig [Module.Finite R M] : Module.Finite R <| AEval R M a := + ‹Module.Finite R M› instance instModulePolynomial : Module R[X] <| AEval R M a := compHom M (aeval a).toRingHom @@ -79,7 +80,7 @@ instance instIsScalarTowerOrigPolynomial : IsScalarTower R R[X] <| AEval R M a w apply (of R M a).symm.injective rw [of_symm_smul, map_smul, smul_assoc, map_smul, of_symm_smul] -instance instFinitePolynomial [Finite R M] : Finite R[X] <| AEval R M a := +instance instFinitePolynomial [Module.Finite R M] : Module.Finite R[X] <| AEval R M a := Finite.of_restrictScalars_finite R _ _ /-- Construct an `R[X]`-linear map out of `AEval R M a` from a `R`-linear map out of `M`. -/ @@ -193,6 +194,6 @@ lemma AEval'.X_smul_of (m : M) : (X : R[X]) • AEval'.of φ m = AEval'.of φ ( lemma AEval'.of_symm_X_smul (m : AEval' φ) : (AEval'.of φ).symm ((X : R[X]) • m) = φ ((AEval'.of φ).symm m) := AEval.of_symm_X_smul _ _ -instance [Finite R M] : Finite R[X] <| AEval' φ := inferInstance +instance [Module.Finite R M] : Module.Finite R[X] <| AEval' φ := inferInstance end Module diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index a063ebfdef089..f78db6c9f424a 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -569,8 +569,8 @@ theorem rank_eq_four [StrongRankCondition R] : Module.rank R ℍ[R,c₁,c₂] = rw [rank_eq_card_basis (basisOneIJK c₁ c₂), Fintype.card_fin] norm_num -theorem finrank_eq_four [StrongRankCondition R] : FiniteDimensional.finrank R ℍ[R,c₁,c₂] = 4 := by - rw [FiniteDimensional.finrank, rank_eq_four, Cardinal.toNat_ofNat] +theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R,c₁,c₂] = 4 := by + rw [Module.finrank, rank_eq_four, Cardinal.toNat_ofNat] /-- There is a natural equivalence when swapping the coefficients of a quaternion algebra. -/ @[simps] @@ -1024,7 +1024,7 @@ instance : Module.Free R ℍ[R] := inferInstanceAs <| Module.Free R ℍ[R,-1,-1] theorem rank_eq_four [StrongRankCondition R] : Module.rank R ℍ[R] = 4 := QuaternionAlgebra.rank_eq_four _ _ -theorem finrank_eq_four [StrongRankCondition R] : FiniteDimensional.finrank R ℍ[R] = 4 := +theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R] = 4 := QuaternionAlgebra.finrank_eq_four _ _ @[simp] theorem star_re : (star a).re = a.re := rfl diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index cbb5edf60ed85..28d32c6a3fea5 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -545,7 +545,7 @@ lemma toClass_eq_zero (P : W.Point) : toClass P = 0 ↔ P = 0 := by rw [← finrank_quotient_span_eq_natDegree_norm (CoordinateRing.basis W) h0, ← (quotientEquivAlgOfEq F hp).toLinearEquiv.finrank_eq, (CoordinateRing.quotientXYIdealEquiv W h).toLinearEquiv.finrank_eq, - FiniteDimensional.finrank_self] + Module.finrank_self] · exact congr_arg toClass lemma toClass_injective : Function.Injective <| @toClass _ _ W := by diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean index 6640b13a2e378..4b70bd94ec863 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean @@ -115,7 +115,7 @@ theorem ae_convolution_tendsto_right_of_locallyIntegrable tendsto_nhdsWithin_iff.2 ⟨hφ, Eventually.of_forall (fun i ↦ (φ i).rOut_pos)⟩ have := (h₀.comp (Besicovitch.tendsto_filterAt μ x₀)).comp hφ' simp only [Function.comp] at this - apply tendsto_integral_smul_of_tendsto_average_norm_sub (K ^ (FiniteDimensional.finrank ℝ G)) this + apply tendsto_integral_smul_of_tendsto_average_norm_sub (K ^ (Module.finrank ℝ G)) this · filter_upwards with i using hg.integrableOn_isCompact (isCompact_closedBall _ _) · apply tendsto_const_nhds.congr (fun i ↦ ?_) diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index a246268057e33..d0f390973978c 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -25,7 +25,7 @@ the indicator function of `closedBall 0 1` with a function as above with `s = ba noncomputable section -open Set Metric TopologicalSpace Function Asymptotics MeasureTheory FiniteDimensional +open Set Metric TopologicalSpace Function Asymptotics MeasureTheory Module ContinuousLinearMap Filter MeasureTheory.Measure Bornology open scoped Pointwise Topology NNReal Convolution diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index 1846ced215ae9..a61f635da6766 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -16,7 +16,7 @@ In this file we define `ContDiffBump.normed f μ` to be the bump function `f` no noncomputable section -open Function Filter Set Metric MeasureTheory FiniteDimensional Measure +open Function Filter Set Metric MeasureTheory Module Measure open scoped Topology namespace ContDiffBump diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index 7b0d3e391e4fd..5b3ed3d9a9e4f 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -24,7 +24,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {D : Type uD} [NormedAddC section FiniteDimensional -open Function FiniteDimensional +open Function Module variable [CompleteSpace 𝕜] diff --git a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean index 6d2876fdc8402..81863eb235bbc 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean @@ -44,7 +44,7 @@ TODO: prove similar theorems assuming that the functions tend to zero at infinit integrable derivatives. -/ -open MeasureTheory Measure FiniteDimensional +open MeasureTheory Measure Module variable {E F G W : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddCommGroup F] [NormedSpace ℝ F] [NormedAddCommGroup G] [NormedSpace ℝ G] [NormedAddCommGroup W] diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index b1924f51dde25..260a55686a38c 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -42,7 +42,7 @@ See `LipschitzWith.hasFderivAt_of_hasLineDerivAt_of_closure`. * [Pertti Mattila, Geometry of sets and measures in Euclidean spaces, Theorem 7.3][Federer1996] -/ -open Filter MeasureTheory Measure FiniteDimensional Metric Set Asymptotics +open Filter MeasureTheory Measure Module Metric Set Asymptotics open scoped NNReal ENNReal Topology diff --git a/Mathlib/Analysis/Convex/Measure.lean b/Mathlib/Analysis/Convex/Measure.lean index d8b9b78be2d4b..cf276ed3f8958 100644 --- a/Mathlib/Analysis/Convex/Measure.lean +++ b/Mathlib/Analysis/Convex/Measure.lean @@ -18,7 +18,7 @@ convex set in `E`. Then the frontier of `s` has measure zero (see `Convex.addHaa open MeasureTheory MeasureTheory.Measure Set Metric Filter Bornology -open FiniteDimensional (finrank) +open Module (finrank) open scoped Topology NNReal ENNReal @@ -64,7 +64,7 @@ theorem addHaar_frontier (hs : Convex ℝ s) : μ (frontier s) = 0 := by /- Due to `Convex.closure_subset_image_homothety_interior_of_one_lt`, for any `r > 1` we have `closure s ⊆ homothety x r '' interior s`, hence `μ (closure s) ≤ r ^ d * μ (interior s)`, where `d = finrank ℝ E`. -/ - set d : ℕ := FiniteDimensional.finrank ℝ E + set d : ℕ := Module.finrank ℝ E have : ∀ r : ℝ≥0, 1 < r → μ (closure s) ≤ ↑(r ^ d) * μ (interior s) := fun r hr ↦ by refine (measure_mono <| hs.closure_subset_image_homothety_interior_of_one_lt hx r hr).trans_eq ?_ diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 2cb4a8b2db0d7..bfb9ac2e29cd3 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -27,7 +27,7 @@ We prove the following facts: variable {ι : Type*} {E P : Type*} -open AffineBasis FiniteDimensional Metric Set +open AffineBasis Module Metric Set open scoped Convex Pointwise Topology section SeminormedAddCommGroup diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index ee505ba41b251..f297649b6f470 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -62,7 +62,7 @@ theorem radon_partition {f : ι → E} (h : ¬ AffineIndependent 𝕜 f) : · linarith only [hI, hJI] · exact (mem_filter.mp hi').2.not_lt (mem_filter.mp hi).2 -open FiniteDimensional +open Module /-- Corner case for `helly_theorem'`. -/ private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 17917e748b4da..e7a99ac85a5d0 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -575,7 +575,7 @@ lemma _root_.ContinuousLinearMap.hasTemperateGrowth (f : E →L[ℝ] F) : variable [NormedAddCommGroup D] [MeasurableSpace D] -open MeasureTheory FiniteDimensional +open MeasureTheory Module /-- A measure `μ` has temperate growth if there is an `n : ℕ` such that `(1 + ‖x‖) ^ (- n)` is `μ`-integrable. -/ @@ -1036,7 +1036,7 @@ section Integration /-! ### Integration -/ -open Real Complex Filter MeasureTheory MeasureTheory.Measure FiniteDimensional +open Real Complex Filter MeasureTheory MeasureTheory.Measure Module variable [RCLike 𝕜] variable [NormedAddCommGroup D] [NormedSpace ℝ D] diff --git a/Mathlib/Analysis/Fourier/Inversion.lean b/Mathlib/Analysis/Fourier/Inversion.lean index cea8c2f5ade12..158aae8391b20 100644 --- a/Mathlib/Analysis/Fourier/Inversion.lean +++ b/Mathlib/Analysis/Fourier/Inversion.lean @@ -37,7 +37,7 @@ To check the concentration property of the middle factor and the fact that it ha rely on the explicit computation of the Fourier transform of Gaussians. -/ -open Filter MeasureTheory Complex FiniteDimensional Metric Real Bornology +open Filter MeasureTheory Complex Module Metric Real Bornology open scoped Topology FourierTransform RealInnerProductSpace Complex diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 5b128034dec9b..1531edb4f4be0 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -45,7 +45,7 @@ equivalence to an inner-product space. noncomputable section -open MeasureTheory Filter Complex Set FiniteDimensional +open MeasureTheory Filter Complex Set Module open scoped Filter Topology Real ENNReal FourierTransform RealInnerProductSpace NNReal diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 6b402f751dabd..5e1e46a1e699c 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -349,7 +349,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] [FiniteDimensional ℝ E] (μ : Measure E) [IsAddHaarMeasure μ] -open FiniteDimensional +open Module /-- The constant factor occurring in the conclusion of `lintegral_pow_le_pow_lintegral_fderiv`. It only depends on `E`, `μ` and `p`. diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 29e41c81cd70d..13a75a6ea3529 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -990,7 +990,7 @@ theorem exists_maximal_orthonormal {s : Set E} (hs : Orthonormal 𝕜 (Subtype.v · exact orthonormal_sUnion_of_directed cc.directedOn fun x xc => hc xc · exact fun _ => Set.subset_sUnion_of_mem -open FiniteDimensional +open Module /-- A family of orthonormal vectors with the correct cardinality forms a basis. -/ def basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι → E} (hv : Orthonormal 𝕜 v) diff --git a/Mathlib/Analysis/InnerProductSpace/EuclideanDist.lean b/Mathlib/Analysis/InnerProductSpace/EuclideanDist.lean index 69889f1ed8e41..50173dd6a64c2 100644 --- a/Mathlib/Analysis/InnerProductSpace/EuclideanDist.lean +++ b/Mathlib/Analysis/InnerProductSpace/EuclideanDist.lean @@ -29,7 +29,7 @@ variable {E : Type*} [AddCommGroup E] [TopologicalSpace E] [TopologicalAddGroup noncomputable section -open FiniteDimensional +open Module /-- If `E` is a finite dimensional space over `ℝ`, then `toEuclidean` is a continuous `ℝ`-linear equivalence between `E` and the Euclidean space of the same dimension. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean b/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean index a3c9277c121f4..b0351cdc69d46 100644 --- a/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean +++ b/Mathlib/Analysis/InnerProductSpace/GramSchmidtOrtho.lean @@ -35,7 +35,7 @@ and outputs a set of orthogonal vectors which have the same span. -/ -open Finset Submodule FiniteDimensional +open Finset Submodule Module variable (𝕜 : Type*) {E : Type*} [RCLike 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] variable {ι : Type*} [LinearOrder ι] [LocallyFiniteOrderBot ι] [IsWellOrder ι (· < ·)] diff --git a/Mathlib/Analysis/InnerProductSpace/Orientation.lean b/Mathlib/Analysis/InnerProductSpace/Orientation.lean index 60b27da658fe3..fb73168b5bab2 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orientation.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orientation.lean @@ -38,7 +38,7 @@ noncomputable section variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] -open FiniteDimensional +open Module open scoped RealInnerProductSpace diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 03d1f51b135c5..5fb4aa1c20293 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -144,11 +144,11 @@ variable [Fintype ι] @[simp] theorem finrank_euclideanSpace : - FiniteDimensional.finrank 𝕜 (EuclideanSpace 𝕜 ι) = Fintype.card ι := by + Module.finrank 𝕜 (EuclideanSpace 𝕜 ι) = Fintype.card ι := by simp [EuclideanSpace, PiLp, WithLp] theorem finrank_euclideanSpace_fin {n : ℕ} : - FiniteDimensional.finrank 𝕜 (EuclideanSpace 𝕜 (Fin n)) = n := by simp + Module.finrank 𝕜 (EuclideanSpace 𝕜 (Fin n)) = n := by simp theorem EuclideanSpace.inner_eq_star_dotProduct (x y : EuclideanSpace 𝕜 ι) : ⟪x, y⟫ = Matrix.dotProduct (star <| WithLp.equiv _ _ x) (WithLp.equiv _ _ y) := @@ -669,7 +669,7 @@ theorem Complex.isometryOfOrthonormal_apply (v : OrthonormalBasis (Fin 2) ℝ F) end Complex -open FiniteDimensional +open Module /-! ### Matrix representation of an orthonormal basis with respect to another -/ @@ -792,7 +792,7 @@ theorem Orthonormal.exists_orthonormalBasis_extension_of_card_eq {ι : Type*} [F obtain ⟨Y, b₀, hX, hb₀⟩ := hX.exists_orthonormalBasis_extension have hιY : Fintype.card ι = Y.card := by refine card_ι.symm.trans ?_ - exact FiniteDimensional.finrank_eq_card_finset_basis b₀.toBasis + exact Module.finrank_eq_card_finset_basis b₀.toBasis have hvsY : s.MapsTo v Y := (s.mapsTo_image v).mono_right (by rwa [← range_restrict]) have hsv' : Set.InjOn v s := by rw [Set.injOn_iff_injective] @@ -840,7 +840,7 @@ irreducible_def DirectSum.IsInternal.sigmaOrthonormalBasisIndexEquiv (hV' : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) : (Σi, Fin (finrank 𝕜 (V i))) ≃ Fin n := let b := hV.collectedOrthonormalBasis hV' fun i => stdOrthonormalBasis 𝕜 (V i) - Fintype.equivFinOfCardEq <| (FiniteDimensional.finrank_eq_card_basis b.toBasis).symm.trans hn + Fintype.equivFinOfCardEq <| (Module.finrank_eq_card_basis b.toBasis).symm.trans hn /-- An `n`-dimensional `InnerProductSpace` equipped with a decomposition as an internal direct sum has an orthonormal basis indexed by `Fin n` and subordinate to that direct sum. -/ @@ -885,7 +885,7 @@ section LinearIsometry variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace 𝕜 V] [FiniteDimensional 𝕜 V] variable {S : Submodule 𝕜 V} {L : S →ₗᵢ[𝕜] V} -open FiniteDimensional +open Module /-- Let `S` be a subspace of a finite-dimensional complex inner product space `V`. A linear isometry mapping `S` into `V` can be extended to a full isometry of `V`. diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index dff8bfa3ecd3d..17223b1faf998 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1029,7 +1029,7 @@ theorem orthogonalProjection_isSymmetric [HasOrthogonalProjection K] : (K.subtypeL ∘L orthogonalProjection K : E →ₗ[𝕜] E).IsSymmetric := inner_orthogonalProjection_left_eq_right K -open FiniteDimensional +open Module /-- Given a finite-dimensional subspace `K₂`, and a subspace `K₁` contained in it, the dimensions of `K₁` and the intersection of its @@ -1038,7 +1038,7 @@ theorem Submodule.finrank_add_inf_finrank_orthogonal {K₁ K₂ : Submodule 𝕜 [FiniteDimensional 𝕜 K₂] (h : K₁ ≤ K₂) : finrank 𝕜 K₁ + finrank 𝕜 (K₁ᗮ ⊓ K₂ : Submodule 𝕜 E) = finrank 𝕜 K₂ := by haveI : FiniteDimensional 𝕜 K₁ := Submodule.finiteDimensional_of_le h - haveI := proper_rclike 𝕜 K₁ + haveI := FiniteDimensional.proper_rclike 𝕜 K₁ have hd := Submodule.finrank_sup_add_finrank_inf_eq K₁ (K₁ᗮ ⊓ K₂) rw [← inf_assoc, (Submodule.orthogonal_disjoint K₁).eq_bot, bot_inf_eq, finrank_bot, Submodule.sup_orthogonal_inf_of_completeSpace h] at hd @@ -1270,7 +1270,7 @@ section OrthonormalBasis variable {v : Set E} -open FiniteDimensional Submodule Set +open Module Submodule Set /-- An orthonormal set in an `InnerProductSpace` is maximal, if and only if the orthogonal complement of its span is empty. -/ @@ -1341,7 +1341,7 @@ variable [FiniteDimensional 𝕜 E] is a basis. -/ theorem maximal_orthonormal_iff_basis_of_finiteDimensional (hv : Orthonormal 𝕜 ((↑) : v → E)) : (∀ u ⊇ v, Orthonormal 𝕜 ((↑) : u → E) → u = v) ↔ ∃ b : Basis v 𝕜 E, ⇑b = ((↑) : v → E) := by - haveI := proper_rclike 𝕜 (span 𝕜 v) + haveI := FiniteDimensional.proper_rclike 𝕜 (span 𝕜 v) rw [maximal_orthonormal_iff_orthogonalComplement_eq_bot hv] rw [Submodule.orthogonal_eq_bot_iff] have hv_coe : range ((↑) : v → E) = v := by simp diff --git a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean index 68d1c7ecfc1b4..0f4bd55b7bae8 100644 --- a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean +++ b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean @@ -184,7 +184,7 @@ end Version1 section Version2 -variable {n : ℕ} (hn : FiniteDimensional.finrank 𝕜 E = n) +variable {n : ℕ} (hn : Module.finrank 𝕜 E = n) /-- A choice of orthonormal basis of eigenvectors for self-adjoint operator `T` on a finite-dimensional inner product space `E`. diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index 1f3cbc87585d7..329dd441e2adc 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -71,7 +71,7 @@ noncomputable section open scoped RealInnerProductSpace ComplexConjugate -open FiniteDimensional +open Module lemma FiniteDimensional.of_fact_finrank_eq_two {K V : Type*} [DivisionRing K] [AddCommGroup V] [Module K V] [Fact (finrank K V = 2)] : FiniteDimensional K V := @@ -204,7 +204,7 @@ def rightAngleRotationAux₂ : E →ₗᵢ[ℝ] E := exact o.areaForm_le x (o.rightAngleRotationAux₁ x) · let K : Submodule ℝ E := ℝ ∙ x have : Nontrivial Kᗮ := by - apply @FiniteDimensional.nontrivial_of_finrank_pos ℝ + apply nontrivial_of_finrank_pos (R := ℝ) have : finrank ℝ K ≤ Finset.card {x} := by rw [← Set.toFinset_singleton] exact finrank_span_le_card ({x} : Set E) diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 4c55f7d09dafa..4735888f210dd 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -48,7 +48,7 @@ universe u v w x noncomputable section -open Set FiniteDimensional TopologicalSpace Filter Asymptotics Topology NNReal Metric +open Asymptotics Filter Module Metric Module NNReal Set TopologicalSpace Topology namespace LinearIsometry @@ -323,14 +323,14 @@ theorem Basis.exists_opNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : instance [FiniteDimensional 𝕜 E] [SecondCountableTopology F] : SecondCountableTopology (E →L[𝕜] F) := by - set d := FiniteDimensional.finrank 𝕜 E + set d := Module.finrank 𝕜 E suffices ∀ ε > (0 : ℝ), ∃ n : (E →L[𝕜] F) → Fin d → ℕ, ∀ f g : E →L[𝕜] F, n f = n g → dist f g ≤ ε from Metric.secondCountable_of_countable_discretization fun ε ε_pos => ⟨Fin d → ℕ, by infer_instance, this ε ε_pos⟩ intro ε ε_pos obtain ⟨u : ℕ → F, hu : DenseRange u⟩ := exists_dense_seq F - let v := FiniteDimensional.finBasis 𝕜 E + let v := Module.finBasis 𝕜 E obtain ⟨C : ℝ, C_pos : 0 < C, hC : ∀ {φ : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ‖φ (v i)‖ ≤ M) → ‖φ‖ ≤ C * M⟩ := @@ -647,7 +647,7 @@ theorem summable_norm_iff {α E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ refine ⟨Summable.of_norm, fun hf ↦ ?_⟩ -- First we use a finite basis to reduce the problem to the case `E = Fin N → ℝ` suffices ∀ {N : ℕ} {g : α → Fin N → ℝ}, Summable g → Summable fun x => ‖g x‖ by - obtain v := finBasis ℝ E + obtain v := Module.finBasis ℝ E set e := v.equivFunL have H : Summable fun x => ‖e (f x)‖ := this (e.summable.2 hf) refine .of_norm_bounded _ (H.mul_left ↑‖(e.symm : (Fin (finrank ℝ E) → ℝ) →L[ℝ] E)‖₊) fun i ↦ ?_ diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean index ce018727e6c0f..b158eee9c860d 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean @@ -107,7 +107,7 @@ theorem exists_extension_norm_eq (p : Subspace 𝕜 E) (f : p →L[𝕜] 𝕜) : _ = ‖f‖ := by rw [reCLM_norm, one_mul] · exact f.opNorm_le_bound g.extendTo𝕜.opNorm_nonneg fun x => h x ▸ g.extendTo𝕜.le_opNorm x -open FiniteDimensional +open Module /-- Corollary of the **Hahn-Banach theorem**: if `f : p → F` is a continuous linear map from a submodule of a normed space `E` over `𝕜`, `𝕜 = ℝ` or `𝕜 = ℂ`, @@ -120,7 +120,7 @@ lemma ContinuousLinearMap.exist_extension_of_finiteDimensional_range {p : Submod (f : p →L[𝕜] F) [FiniteDimensional 𝕜 (LinearMap.range f)] : ∃ g : E →L[𝕜] F, f = g.comp p.subtypeL := by letI : RCLike 𝕜 := IsRCLikeNormedField.rclike 𝕜 - set b := finBasis 𝕜 (LinearMap.range f) + set b := Module.finBasis 𝕜 (LinearMap.range f) set e := b.equivFunL set fi := fun i ↦ (LinearMap.toContinuousLinearMap (b.coord i)).comp (f.codRestrict _ <| LinearMap.mem_range_self _) diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean index ea572b4ce3c4e..89db8290e3687 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean @@ -332,18 +332,18 @@ theorem integral_cexp_neg_mul_sq_norm_add_of_euclideanSpace theorem integral_cexp_neg_mul_sq_norm_add (hb : 0 < b.re) (c : ℂ) (w : V) : ∫ v : V, cexp (- b * ‖v‖^2 + c * ⟪w, v⟫) = - (π / b) ^ (FiniteDimensional.finrank ℝ V / 2 : ℂ) * cexp (c ^ 2 * ‖w‖^2 / (4 * b)) := by + (π / b) ^ (Module.finrank ℝ V / 2 : ℂ) * cexp (c ^ 2 * ‖w‖^2 / (4 * b)) := by let e := (stdOrthonormalBasis ℝ V).repr.symm rw [← e.measurePreserving.integral_comp e.toHomeomorph.measurableEmbedding] convert integral_cexp_neg_mul_sq_norm_add_of_euclideanSpace hb c (e.symm w) <;> simp [LinearIsometryEquiv.inner_map_eq_flip] theorem integral_cexp_neg_mul_sq_norm (hb : 0 < b.re) : - ∫ v : V, cexp (- b * ‖v‖^2) = (π / b) ^ (FiniteDimensional.finrank ℝ V / 2 : ℂ) := by + ∫ v : V, cexp (- b * ‖v‖^2) = (π / b) ^ (Module.finrank ℝ V / 2 : ℂ) := by simpa using integral_cexp_neg_mul_sq_norm_add hb 0 (0 : V) theorem integral_rexp_neg_mul_sq_norm {b : ℝ} (hb : 0 < b) : - ∫ v : V, rexp (- b * ‖v‖^2) = (π / b) ^ (FiniteDimensional.finrank ℝ V / 2 : ℝ) := by + ∫ v : V, rexp (- b * ‖v‖^2) = (π / b) ^ (Module.finrank ℝ V / 2 : ℝ) := by rw [← ofReal_inj] convert integral_cexp_neg_mul_sq_norm (show 0 < (b : ℂ).re from hb) (V := V) · change ofRealLI (∫ (v : V), rexp (-b * ‖v‖ ^ 2)) = ∫ (v : V), cexp (-↑b * ↑‖v‖ ^ 2) @@ -354,7 +354,7 @@ theorem integral_rexp_neg_mul_sq_norm {b : ℝ} (hb : 0 < b) : theorem _root_.fourierIntegral_gaussian_innerProductSpace' (hb : 0 < b.re) (x w : V) : 𝓕 (fun v ↦ cexp (- b * ‖v‖^2 + 2 * π * Complex.I * ⟪x, v⟫)) w = - (π / b) ^ (FiniteDimensional.finrank ℝ V / 2 : ℂ) * cexp (-π ^ 2 * ‖x - w‖ ^ 2 / b) := by + (π / b) ^ (Module.finrank ℝ V / 2 : ℂ) * cexp (-π ^ 2 * ‖x - w‖ ^ 2 / b) := by simp only [neg_mul, fourierIntegral_eq', ofReal_neg, ofReal_mul, ofReal_ofNat, smul_eq_mul, ← Complex.exp_add, real_inner_comm w] convert integral_cexp_neg_mul_sq_norm_add hb (2 * π * Complex.I) (x - w) using 3 with v @@ -367,7 +367,7 @@ theorem _root_.fourierIntegral_gaussian_innerProductSpace' (hb : 0 < b.re) (x w theorem _root_.fourierIntegral_gaussian_innerProductSpace (hb : 0 < b.re) (w : V) : 𝓕 (fun v ↦ cexp (- b * ‖v‖^2)) w = - (π / b) ^ (FiniteDimensional.finrank ℝ V / 2 : ℂ) * cexp (-π ^ 2 * ‖w‖^2 / b) := by + (π / b) ^ (Module.finrank ℝ V / 2 : ℂ) * cexp (-π ^ 2 * ‖w‖^2 / b) := by simpa using fourierIntegral_gaussian_innerProductSpace' hb 0 w end InnerProductSpace diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 86a45e8436078..5b9cbab201ab2 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -27,7 +27,7 @@ noncomputable section open scoped NNReal Filter Topology ENNReal -open Asymptotics Filter Set Real MeasureTheory FiniteDimensional +open Asymptotics Filter Set Real MeasureTheory Module variable {E : Type*} [NormedAddCommGroup E] diff --git a/Mathlib/CategoryTheory/Preadditive/Schur.lean b/Mathlib/CategoryTheory/Preadditive/Schur.lean index 407e246e350e0..2fd3b980f6f59 100644 --- a/Mathlib/CategoryTheory/Preadditive/Schur.lean +++ b/Mathlib/CategoryTheory/Preadditive/Schur.lean @@ -70,7 +70,7 @@ noncomputable instance [HasKernels C] {X : C} [Simple X] : DivisionRing (End X) qsmul := _ qsmul_def := fun q a => rfl -open FiniteDimensional +open Module section diff --git a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean index 40f0ebf38db1e..e609a7dee3f0e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean @@ -189,8 +189,8 @@ end /-- The number of connected components in `G` is the dimension of the nullspace its Laplacian. -/ theorem card_ConnectedComponent_eq_rank_ker_lapMatrix : Fintype.card G.ConnectedComponent = - FiniteDimensional.finrank ℝ (LinearMap.ker (Matrix.toLin' (G.lapMatrix ℝ))) := by + Module.finrank ℝ (LinearMap.ker (Matrix.toLin' (G.lapMatrix ℝ))) := by classical - rw [FiniteDimensional.finrank_eq_card_basis (lapMatrix_ker_basis G)] + rw [Module.finrank_eq_card_basis (lapMatrix_ker_basis G)] end SimpleGraph diff --git a/Mathlib/Data/Complex/FiniteDimensional.lean b/Mathlib/Data/Complex/FiniteDimensional.lean index 5fbde35674c53..f989ee2a1847c 100644 --- a/Mathlib/Data/Complex/FiniteDimensional.lean +++ b/Mathlib/Data/Complex/FiniteDimensional.lean @@ -15,12 +15,11 @@ This file contains the `FiniteDimensional ℝ ℂ` instance, as well as some res (`finrank` and `Module.rank`). -/ -open FiniteDimensional +open Module namespace Complex -instance : FiniteDimensional ℝ ℂ := - of_fintype_basis basisOneI +instance : FiniteDimensional ℝ ℂ := .of_fintype_basis basisOneI @[simp] theorem finrank_real_complex : finrank ℝ ℂ = 2 := by @@ -50,8 +49,8 @@ theorem rank_real_of_complex (E : Type*) [AddCommGroup E] [Module ℂ E] : simp only [Cardinal.lift_id'] theorem finrank_real_of_complex (E : Type*) [AddCommGroup E] [Module ℂ E] : - FiniteDimensional.finrank ℝ E = 2 * FiniteDimensional.finrank ℂ E := by - rw [← FiniteDimensional.finrank_mul_finrank ℝ ℂ E, Complex.finrank_real_complex] + Module.finrank ℝ E = 2 * Module.finrank ℂ E := by + rw [← Module.finrank_mul_finrank ℝ ℂ E, Complex.finrank_real_complex] section Rational diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index d272ff6eb228a..13db0d03e2bc9 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -26,7 +26,7 @@ open Matrix namespace Matrix -open FiniteDimensional +open Module variable {l m n o R : Type*} [Fintype n] [Fintype o] @@ -168,7 +168,7 @@ variable [Field R] /-- The rank of a diagnonal matrix is the count of non-zero elements on its main diagonal -/ theorem rank_diagonal [Fintype m] [DecidableEq m] [DecidableEq R] (w : m → R) : (diagonal w).rank = Fintype.card {i // (w i) ≠ 0} := by - rw [Matrix.rank, ← Matrix.toLin'_apply', FiniteDimensional.finrank, ← LinearMap.rank, + rw [Matrix.rank, ← Matrix.toLin'_apply', Module.finrank, ← LinearMap.rank, LinearMap.rank_diagonal, Cardinal.toNat_natCast] end Field @@ -278,7 +278,7 @@ lemma rank_add_rank_le_card_of_mul_eq_zero [Field R] [Finite l] [Fintype m] let en : Basis n R (n → R) := Pi.basisFun R n rw [Matrix.rank_eq_finrank_range_toLin A el em, Matrix.rank_eq_finrank_range_toLin B em en, - ← FiniteDimensional.finrank_fintype_fun_eq_card R, + ← Module.finrank_fintype_fun_eq_card R, ← LinearMap.finrank_range_add_finrank_ker (Matrix.toLin em el A), add_le_add_iff_left] apply Submodule.finrank_mono diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index e4701e89934ea..c83073349810a 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -28,7 +28,7 @@ For example, `Algebra.adjoin K {x}` might not include `x⁻¹`. - `F⟮α⟯`: adjoin a single element `α` to `F` (in scope `IntermediateField`). -/ -open FiniteDimensional Polynomial +open Module Polynomial namespace IntermediateField @@ -872,7 +872,7 @@ theorem adjoin_natCast (n : ℕ) : F⟮(n : E)⟯ = ⊥ := section AdjoinRank -open FiniteDimensional Module +open Module Module variable {K L : IntermediateField F E} @@ -1043,7 +1043,7 @@ theorem isAlgebraic_adjoin_simple {x : L} (hx : IsIntegral K x) : Algebra.IsAlge have := adjoin.finiteDimensional hx; Algebra.IsAlgebraic.of_finite K K⟮x⟯ theorem adjoin.finrank {x : L} (hx : IsIntegral K x) : - FiniteDimensional.finrank K K⟮x⟯ = (minpoly K x).natDegree := by + Module.finrank K K⟮x⟯ = (minpoly K x).natDegree := by rw [PowerBasis.finrank (adjoin.powerBasis hx : _)] rfl @@ -1117,11 +1117,10 @@ theorem _root_.minpoly.degree_le (x : L) [FiniteDimensional K L] : /-- If `x : L` is an integral element in a field extension `L` over `K`, then the degree of the minimal polynomial of `x` over `K` divides `[L : K]`.-/ theorem _root_.minpoly.degree_dvd {x : L} (hx : IsIntegral K x) : - (minpoly K x).natDegree ∣ FiniteDimensional.finrank K L := by + (minpoly K x).natDegree ∣ finrank K L := by rw [dvd_iff_exists_eq_mul_left, ← IntermediateField.adjoin.finrank hx] - use FiniteDimensional.finrank K⟮x⟯ L - rw [eq_comm, mul_comm] - exact FiniteDimensional.finrank_mul_finrank _ _ _ + use finrank K⟮x⟯ L + rw [mul_comm, finrank_mul_finrank] -- TODO: generalize to `Sort` /-- A compositum of algebraic extensions is algebraic -/ @@ -1178,7 +1177,7 @@ theorem card_algHom_adjoin_integral (h : IsIntegral F α) (h_sep : IsSeparable F exact h_sep -- Apparently `K⟮root f⟯ →+* K⟮root f⟯` is expensive to unify during instance synthesis. -open FiniteDimensional AdjoinRoot in +open Module AdjoinRoot in /-- Let `f, g` be monic polynomials over `K`. If `f` is irreducible, and `g(x) - α` is irreducible in `K⟮α⟯` with `α` a root of `f`, then `f(g(x))` is irreducible. -/ theorem _root_.Polynomial.irreducible_comp {f g : K[X]} (hfm : f.Monic) (hgm : g.Monic) @@ -1226,7 +1225,7 @@ theorem _root_.Polynomial.irreducible_comp {f g : K[X]} (hfm : f.Monic) (hgm : g rw [← finrank_top', ← this, adjoin.finrank] exact IsIntegral.of_finite _ _ · simp [← key₂] - have := FiniteDimensional.finrank_mul_finrank K K⟮aeval (root p) g⟯ Kx + have := Module.finrank_mul_finrank K K⟮aeval (root p) g⟯ Kx rwa [key₁', key₂', (AdjoinRoot.powerBasis hp₁.ne_zero).finrank, powerBasis_dim, eq_comm] at this end AdjoinIntegralElement diff --git a/Mathlib/FieldTheory/Cardinality.lean b/Mathlib/FieldTheory/Cardinality.lean index 8a0482682912f..045db9322dcbf 100644 --- a/Mathlib/FieldTheory/Cardinality.lean +++ b/Mathlib/FieldTheory/Cardinality.lean @@ -43,8 +43,8 @@ theorem Fintype.isPrimePow_card_of_field {α} [Fintype α] [Field α] : IsPrimeP let b := IsNoetherian.finsetBasis (ZMod p) α rw [Module.card_fintype b, ZMod.card, isPrimePow_pow_iff] · exact hp.1.isPrimePow - rw [← FiniteDimensional.finrank_eq_card_basis b] - exact FiniteDimensional.finrank_pos.ne' + rw [← Module.finrank_eq_card_basis b] + exact Module.finrank_pos.ne' /-- A `Fintype` can be given a field structure iff its cardinality is a prime power. -/ theorem Fintype.nonempty_field_iff {α} [Fintype α] : Nonempty (Field α) ↔ IsPrimePow ‖α‖ := by diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 6b5c4835a87c6..715387ba0669b 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -455,9 +455,9 @@ variable {V : Type*} [Fintype K] [DivisionRing K] [AddCommGroup V] [Module K V] -- should this go in a namespace? -- finite_dimensional would be natural, -- but we don't assume it... -theorem card_eq_pow_finrank [Fintype V] : Fintype.card V = q ^ FiniteDimensional.finrank K V := by +theorem card_eq_pow_finrank [Fintype V] : Fintype.card V = q ^ Module.finrank K V := by let b := IsNoetherian.finsetBasis K V - rw [Module.card_fintype b, ← FiniteDimensional.finrank_eq_card_basis b] + rw [Module.card_fintype b, ← Module.finrank_eq_card_basis b] end diff --git a/Mathlib/FieldTheory/Finite/GaloisField.lean b/Mathlib/FieldTheory/Finite/GaloisField.lean index e03d1331f1020..89b627d0328c8 100644 --- a/Mathlib/FieldTheory/Finite/GaloisField.lean +++ b/Mathlib/FieldTheory/Finite/GaloisField.lean @@ -88,7 +88,7 @@ instance : Fintype (GaloisField p n) := by dsimp only [GaloisField] exact FiniteDimensional.fintypeOfFintype (ZMod p) (GaloisField p n) -theorem finrank {n} (h : n ≠ 0) : FiniteDimensional.finrank (ZMod p) (GaloisField p n) = n := by +theorem finrank {n} (h : n ≠ 0) : Module.finrank (ZMod p) (GaloisField p n) = n := by set g_poly := (X ^ p ^ n - X : (ZMod p)[X]) have hp : 1 < p := h_prime.out.one_lt have aux : g_poly ≠ 0 := FiniteField.X_pow_card_pow_sub_X_ne_zero _ h hp @@ -139,7 +139,7 @@ theorem finrank {n} (h : n ≠ 0) : FiniteDimensional.finrank (ZMod p) (GaloisFi theorem card (h : n ≠ 0) : Fintype.card (GaloisField p n) = p ^ n := by let b := IsNoetherian.finsetBasis (ZMod p) (GaloisField p n) - rw [Module.card_fintype b, ← FiniteDimensional.finrank_eq_card_basis b, ZMod.card, finrank p h] + rw [Module.card_fintype b, ← Module.finrank_eq_card_basis b, ZMod.card, finrank p h] theorem splits_zmod_X_pow_sub_X : Splits (RingHom.id (ZMod p)) (X ^ p - X) := by have hp : 1 < p := h_prime.out.one_lt diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index 8207de7d06c14..00e92545eba5b 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -215,8 +215,8 @@ instance [Finite σ] : FiniteDimensional K (R σ K) := by simpa only [rank_R] using Cardinal.nat_lt_aleph0 (Fintype.card (σ → K))) open Classical in -theorem finrank_R [Fintype σ] : FiniteDimensional.finrank K (R σ K) = Fintype.card (σ → K) := - FiniteDimensional.finrank_eq_of_rank_eq (rank_R σ K) +theorem finrank_R [Fintype σ] : Module.finrank K (R σ K) = Fintype.card (σ → K) := + Module.finrank_eq_of_rank_eq (rank_R σ K) -- Porting note: was `(evalᵢ σ K).range`. theorem range_evalᵢ [Finite σ] : range (evalᵢ σ K) = ⊤ := by @@ -228,7 +228,7 @@ theorem ker_evalₗ [Finite σ] : ker (evalᵢ σ K) = ⊥ := by cases nonempty_fintype σ refine (ker_eq_bot_iff_range_eq_top_of_finrank_eq_finrank ?_).mpr (range_evalᵢ σ K) classical - rw [FiniteDimensional.finrank_fintype_fun_eq_card, finrank_R] + rw [Module.finrank_fintype_fun_eq_card, finrank_R] theorem eq_zero_of_eval_eq_zero [Finite σ] (p : MvPolynomial σ K) (h : ∀ v : σ → K, eval v p = 0) (hp : p ∈ restrictDegree σ K (Fintype.card K - 1)) : p = 0 := diff --git a/Mathlib/FieldTheory/Fixed.lean b/Mathlib/FieldTheory/Fixed.lean index bf86498f3027a..965f23a35eefe 100644 --- a/Mathlib/FieldTheory/Fixed.lean +++ b/Mathlib/FieldTheory/Fixed.lean @@ -30,7 +30,7 @@ element of `G`, where `G` is a group that acts on `F`. noncomputable section -open MulAction Finset FiniteDimensional +open MulAction Finset Module universe u v w diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 56ce629fc2bea..3f571ef51fa06 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -35,7 +35,7 @@ Together, these two results prove the Galois correspondence. open scoped Polynomial IntermediateField -open FiniteDimensional AlgEquiv +open Module AlgEquiv section diff --git a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean index 99eab922d6a3d..55ecc96148306 100644 --- a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean @@ -13,7 +13,7 @@ import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition # Results on finite dimensionality and algebraicity of intermediate fields. -/ -open FiniteDimensional +open Module variable {K : Type*} {L : Type*} [Field K] [Field L] [Algebra K L] {S : IntermediateField K L} @@ -24,11 +24,8 @@ section FiniteDimensional variable (F E : IntermediateField K L) -instance finiteDimensional_left [FiniteDimensional K L] : FiniteDimensional K F := - left K F L - -instance finiteDimensional_right [FiniteDimensional K L] : FiniteDimensional F L := - right K F L +instance finiteDimensional_left [FiniteDimensional K L] : FiniteDimensional K F := .left K F L +instance finiteDimensional_right [FiniteDimensional K L] : FiniteDimensional F L := .right K F L @[simp] theorem rank_eq_rank_subalgebra : Module.rank K F.toSubalgebra = Module.rank K F := diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index 70d3708472eb5..743199124c340 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -60,7 +60,7 @@ perfect ring, perfect closure, purely inseparable -/ -open FiniteDimensional Polynomial IntermediateField Field +open Module Polynomial IntermediateField Field noncomputable section diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 4be37a02bc718..939bc161ec3f8 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -530,7 +530,7 @@ lemma isGalois_of_isSplittingField_X_pow_sub_C : IsGalois K L := IsGalois.of_separable_splitting_field (separable_X_pow_sub_C_of_irreducible hζ a H) include hζ H in -lemma finrank_of_isSplittingField_X_pow_sub_C : FiniteDimensional.finrank K L = n := by +lemma finrank_of_isSplittingField_X_pow_sub_C : Module.finrank K L = n := by have := Polynomial.IsSplittingField.finiteDimensional L (X ^ n - C a) have := isGalois_of_isSplittingField_X_pow_sub_C hζ H L have hn := Nat.pos_iff_ne_zero.mpr (ne_zero_of_irreducible_X_pow_sub_C H) @@ -545,9 +545,9 @@ end IsSplittingField section IsCyclic variable {L} [Field L] [Algebra K L] [FiniteDimensional K L] -variable (hK : (primitiveRoots (FiniteDimensional.finrank K L) K).Nonempty) +variable (hK : (primitiveRoots (Module.finrank K L) K).Nonempty) -open FiniteDimensional +open Module variable (K L) include hK in @@ -623,7 +623,7 @@ lemma isSplittingField_X_pow_sub_C_of_root_adjoin_eq_top end IsCyclic -open FiniteDimensional in +open Module in /-- Suppose `L/K` is a finite extension of dimension `n`, and `K` contains all `n`-th roots of unity. Then `L/K` is cyclic iff @@ -631,7 +631,7 @@ Then `L/K` is cyclic iff `L = K[α]` for some `αⁿ ∈ K`. -/ lemma isCyclic_tfae (K L) [Field K] [Field L] [Algebra K L] [FiniteDimensional K L] - (hK : (primitiveRoots (FiniteDimensional.finrank K L) K).Nonempty) : + (hK : (primitiveRoots (Module.finrank K L) K).Nonempty) : List.TFAE [ IsGalois K L ∧ IsCyclic (L ≃ₐ[K] L), ∃ a : K, Irreducible (X ^ (finrank K L) - C a) ∧ diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index b62efc1339b8a..3a33353b9af98 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -162,7 +162,7 @@ variable (F E K : Type*) [Field F] [Ring E] [CommRing K] [IsDomain K] [Algebra F -- 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 (FiniteDimensional.finBasis F E : _ → E)) : + (x : range (Module.finBasis F E : _ → E)) : { l : K // l ∈ (minpoly F x.1).aroots K } := ⟨φ x, by rw [mem_roots_map (minpoly.ne_zero_of_finite F x.val), @@ -173,14 +173,14 @@ theorem aux_inj_roots_of_min_poly : Injective (rootsOfMinPolyPiType F E K) := by -- needs explicit coercion on the RHS suffices (f : E →ₗ[F] K) = (g : E →ₗ[F] K) by rwa [DFunLike.ext'_iff] at this ⊢ rw [funext_iff] at h - exact LinearMap.ext_on (FiniteDimensional.finBasis F E).span_eq fun e he => + exact LinearMap.ext_on (Module.finBasis F E).span_eq fun e he => Subtype.ext_iff.mp (h ⟨e, he⟩) /-- Given field extensions `E/F` and `K/F`, with `E/F` finite, there are finitely many `F`-algebra homomorphisms `E →ₐ[K] K`. -/ noncomputable instance AlgHom.fintype : Fintype (E →ₐ[F] K) := @Fintype.ofInjective _ _ - (Fintype.subtypeProd (finite_range (FiniteDimensional.finBasis F E)) fun e => + (Fintype.subtypeProd (finite_range (Module.finBasis F E)) fun e => (minpoly F e).aroots K) _ (aux_inj_roots_of_min_poly F E K) diff --git a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean index c327e135c3ddd..fae1bec1703c0 100644 --- a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean +++ b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean @@ -20,7 +20,7 @@ See `traceForm_dualBasis_powerBasis_eq`. - `span_coeff_minpolyDiv`: The coefficients of `minpolyDiv` spans `R`. -/ -open Polynomial FiniteDimensional +open Polynomial Module variable (R K) {L S} [CommRing R] [Field K] [Field L] [CommRing S] [Algebra R S] [Algebra K L] variable (x : S) diff --git a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean index 003d87b0e80d0..ef2d383b8d52b 100644 --- a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean +++ b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean @@ -41,7 +41,7 @@ noncomputable section open scoped Polynomial -open FiniteDimensional +open Module namespace Polynomial @@ -384,10 +384,10 @@ theorem prime_degree_dvd_card [CharZero F] (p_irr : Irreducible p) (p_deg : p.na let α : p.SplittingField := rootOfSplits (algebraMap F p.SplittingField) (SplittingField.splits p) hp have hα : IsIntegral F α := .of_finite F α - use FiniteDimensional.finrank F⟮α⟯ p.SplittingField + use Module.finrank F⟮α⟯ p.SplittingField suffices (minpoly F α).natDegree = p.natDegree by letI _ : AddCommGroup F⟮α⟯ := Ring.toAddCommGroup - rw [← FiniteDimensional.finrank_mul_finrank F F⟮α⟯ p.SplittingField, + rw [← Module.finrank_mul_finrank F F⟮α⟯ p.SplittingField, IntermediateField.adjoin.finrank hα, this] suffices minpoly F α ∣ p by have key := (minpoly.irreducible hα).dvd_symm p_irr this diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 65b6d1707f19d..da215b5d8dbf4 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -36,7 +36,7 @@ exists_adjoin_simple_eq_top noncomputable section -open FiniteDimensional Polynomial IntermediateField +open Module Polynomial IntermediateField namespace Field @@ -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 := finite_of_finite F E + haveI : Finite E := FiniteDimensional.finite_of_finite F E exists_primitive_element_of_finite_top F E end PrimitiveElementFinite @@ -367,7 +367,7 @@ section iff namespace Field -open FiniteDimensional IntermediateField Polynomial Algebra Set +open Module IntermediateField Polynomial Algebra Set variable (F : Type*) {E : Type*} [Field F] [Field E] [Algebra F E] [FiniteDimensional F E] diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index f8b347c9bf68a..c880a81782d7f 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -127,7 +127,7 @@ separable degree, degree, separable closure, purely inseparable -/ -open FiniteDimensional Polynomial IntermediateField Field Finsupp +open Module Polynomial IntermediateField Field Finsupp noncomputable section @@ -756,7 +756,7 @@ private theorem LinearIndependent.map_pow_expChar_pow_of_fd_isSeparable have h' := h.coe_range let ι' := h'.extend (Set.range v).subset_univ let b : Basis ι' F E := Basis.extend h' - letI : Fintype ι' := fintypeBasisIndex b + letI : Fintype ι' := FiniteDimensional.fintypeBasisIndex b have H := linearIndependent_of_top_le_span_of_card_eq_finrank (span_map_pow_expChar_pow_eq_top_of_isSeparable q n b.span_eq).ge (finrank_eq_card_basis b).symm diff --git a/Mathlib/FieldTheory/SeparableClosure.lean b/Mathlib/FieldTheory/SeparableClosure.lean index 596db72fdf6cd..68110800b109d 100644 --- a/Mathlib/FieldTheory/SeparableClosure.lean +++ b/Mathlib/FieldTheory/SeparableClosure.lean @@ -59,7 +59,7 @@ separable degree, degree, separable closure -/ -open FiniteDimensional Polynomial IntermediateField Field +open Module Polynomial IntermediateField Field noncomputable section diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 3bac952cdb9f8..6ebcf74215644 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -67,7 +67,7 @@ This file contains basics about the separable degree of a field extension. if `K / E / F` is a field extension tower, such that `K / E` is algebraic, then there is a non-canonical bijection `Field.Emb F E × Field.Emb E K ≃ Field.Emb F K`. In particular, the separable degrees satisfy the tower law: $[E:F]_s [K:E]_s = [K:F]_s$ - (see also `FiniteDimensional.finrank_mul_finrank`). + (see also `Module.finrank_mul_finrank`). - `Polynomial.natSepDegree_le_natDegree`: the separable degree of a polynomial is smaller than its degree. @@ -118,7 +118,7 @@ separable degree, degree, polynomial -/ -open FiniteDimensional Polynomial IntermediateField Field +open Module Polynomial IntermediateField Field noncomputable section @@ -246,7 +246,7 @@ def embProdEmbOfIsAlgebraic [Algebra E K] [IsScalarTower F E K] [Algebra.IsAlgeb /-- If `K / E / F` is a field extension tower, such that `K / E` is algebraic, then their separable degrees satisfy the tower law -$[E:F]_s [K:E]_s = [K:F]_s$. See also `FiniteDimensional.finrank_mul_finrank`. -/ +$[E:F]_s [K:E]_s = [K:F]_s$. See also `Module.finrank_mul_finrank`. -/ theorem finSepDegree_mul_finSepDegree_of_isAlgebraic [Algebra E K] [IsScalarTower F E K] [Algebra.IsAlgebraic E K] : finSepDegree F E * finSepDegree E K = finSepDegree F K := by @@ -701,7 +701,7 @@ theorem finSepDegree_dvd_finrank : finSepDegree F E ∣ finrank F E := by set M := L⟮x⟯ have := Algebra.IsAlgebraic.of_finite L M rwa [finSepDegree_mul_finSepDegree_of_isAlgebraic F L M, - FiniteDimensional.finrank_mul_finrank F L M] at hdvd + Module.finrank_mul_finrank F L M] at hdvd rw [finrank_of_infinite_dimensional hfd] exact dvd_zero _ @@ -735,7 +735,7 @@ theorem finSepDegree_eq_finrank_of_isSeparable [Algebra.IsSeparable F E] : set M := L⟮x⟯ have := Algebra.IsAlgebraic.of_finite L M rwa [finSepDegree_mul_finSepDegree_of_isAlgebraic F L M, - FiniteDimensional.finrank_mul_finrank F L M] at heq + Module.finrank_mul_finrank F L M] at heq alias Algebra.IsSeparable.finSepDegree_eq := finSepDegree_eq_finrank_of_isSeparable @@ -749,7 +749,7 @@ theorem finSepDegree_eq_finrank_iff [FiniteDimensional F E] : (finSepDegree_adjoin_simple_le_finrank F E x halg) <| le_of_not_lt fun h ↦ ?_ have := Nat.mul_lt_mul_of_lt_of_le' h (finSepDegree_le_finrank F⟮x⟯ E) Fin.size_pos' rw [finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E, - FiniteDimensional.finrank_mul_finrank F F⟮x⟯ E] at this + Module.finrank_mul_finrank F F⟮x⟯ E] at this linarith only [heq, this]⟩, fun _ ↦ finSepDegree_eq_finrank_of_isSeparable F E⟩ end Field @@ -796,7 +796,7 @@ theorem IsSeparable.of_algebra_isSeparable_of_isSeparable [Algebra E K] [IsScala have := finSepDegree_mul_finSepDegree_of_isAlgebraic F E' E'⟮x⟯ rw [finSepDegree_eq_finrank_of_isSeparable F E', finSepDegree_eq_finrank_of_isSeparable E' E'⟮x⟯, - FiniteDimensional.finrank_mul_finrank F E' E'⟮x⟯, + Module.finrank_mul_finrank F E' E'⟮x⟯, eq_comm, finSepDegree_eq_finrank_iff F E'⟮x⟯] at this change Algebra.IsSeparable F (restrictScalars F E'⟮x⟯) at this exact isSeparable_of_mem_isSeparable F K hx diff --git a/Mathlib/FieldTheory/Tower.lean b/Mathlib/FieldTheory/Tower.lean index 4a4720b3cb6aa..297650824b8be 100644 --- a/Mathlib/FieldTheory/Tower.lean +++ b/Mathlib/FieldTheory/Tower.lean @@ -15,7 +15,7 @@ We prove that given `IsScalarTower F K A`, if `A` is finite as a module over `F` In particular these conditions hold when `A`, `F`, and `K` are fields. -The formulas for the dimensions are given elsewhere by `FiniteDimensional.finrank_mul_finrank`. +The formulas for the dimensions are given elsewhere by `Module.finrank_mul_finrank`. ## Tags diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean index 40018bcc026e4..6e6e991b9b0b4 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean @@ -22,7 +22,7 @@ This file defines oriented angles in Euclidean affine spaces. noncomputable section -open FiniteDimensional Complex +open Module Complex open scoped Affine EuclideanGeometry Real RealInnerProductSpace ComplexConjugate diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index fb641332e085a..3f2938cdce408 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -31,7 +31,7 @@ modulo `2 * π` as equalities of `(2 : ℤ) • θ`. noncomputable section -open FiniteDimensional Complex +open Module Complex open scoped Real RealInnerProductSpace ComplexConjugate diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/RightAngle.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/RightAngle.lean index d634919578060..8ca1f21c9b812 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/RightAngle.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/RightAngle.lean @@ -25,7 +25,7 @@ open scoped RealInnerProductSpace namespace Orientation -open FiniteDimensional +open Module variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] variable [hd2 : Fact (finrank ℝ V = 2)] (o : Orientation ℝ V (Fin 2)) @@ -519,7 +519,7 @@ end Orientation namespace EuclideanGeometry -open FiniteDimensional +open Module variable {V : Type*} {P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (finrank ℝ V = 2)] [Module.Oriented ℝ V (Fin 2)] diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean index 7f40cc129de8e..734899c867f50 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean @@ -20,7 +20,7 @@ This file defines rotations by oriented angles in real inner product spaces. noncomputable section -open FiniteDimensional Complex +open Module Complex open scoped Real RealInnerProductSpace ComplexConjugate @@ -101,8 +101,7 @@ theorem rotation_eq_matrix_toLin (θ : Real.Angle) {x : V} (hx : x ≠ 0) : /-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ @[simp] theorem det_rotation (θ : Real.Angle) : LinearMap.det (o.rotation θ).toLinearMap = 1 := by - haveI : Nontrivial V := - FiniteDimensional.nontrivial_of_finrank_eq_succ (@Fact.out (finrank ℝ V = 2) _) + haveI : Nontrivial V := nontrivial_of_finrank_eq_succ (@Fact.out (finrank ℝ V = 2) _) obtain ⟨x, hx⟩ : ∃ x, x ≠ (0 : V) := exists_ne (0 : V) rw [o.rotation_eq_matrix_toLin θ hx] simpa [sq] using θ.cos_sq_add_sin_sq @@ -333,8 +332,7 @@ theorem oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : Real.Angle theorem exists_linearIsometryEquiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} (hd : 0 < LinearMap.det (f.toLinearEquiv : V →ₗ[ℝ] V)) : ∃ θ : Real.Angle, f = o.rotation θ := by - haveI : Nontrivial V := - FiniteDimensional.nontrivial_of_finrank_eq_succ (@Fact.out (finrank ℝ V = 2) _) + haveI : Nontrivial V := nontrivial_of_finrank_eq_succ (@Fact.out (finrank ℝ V = 2) _) obtain ⟨x, hx⟩ : ∃ x, x ≠ (0 : V) := exists_ne (0 : V) use o.oangle x (f x) apply LinearIsometryEquiv.toLinearEquiv_injective diff --git a/Mathlib/Geometry/Euclidean/Angle/Sphere.lean b/Mathlib/Geometry/Euclidean/Angle/Sphere.lean index dd2a20922dbdb..17315a1b24430 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Sphere.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Sphere.lean @@ -16,7 +16,7 @@ This file proves results about angles in circles and spheres. noncomputable section -open FiniteDimensional Complex +open Module Complex open scoped EuclideanGeometry Real RealInnerProductSpace ComplexConjugate diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index 3ac078670754f..55b07f87b916c 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -124,7 +124,7 @@ theorem dist_smul_vadd_eq_dist {v : V} (p₁ p₂ : P) (hv : v ≠ 0) (r : ℝ) mul_div_assoc] norm_num -open AffineSubspace FiniteDimensional +open AffineSubspace Module /-- Distances `r₁` `r₂` of `p` from two different points `c₁` `c₂` determine at most two points `p₁` `p₂` in a two-dimensional subspace containing those points @@ -150,7 +150,7 @@ theorem eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two {s : AffineSubspace · rw [real_inner_comm] exact ho have hbs : Submodule.span ℝ (Set.range b) = s.direction := by - refine eq_of_le_of_finrank_eq ?_ ?_ + refine Submodule.eq_of_le_of_finrank_eq ?_ ?_ · rw [Submodule.span_le, Set.range_subset_iff] intro i fin_cases i diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index ed19fa6831d20..ab715d6d83899 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -667,7 +667,7 @@ end Affine namespace EuclideanGeometry -open Affine AffineSubspace FiniteDimensional +open Affine AffineSubspace Module variable {V : Type*} {P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] [NormedAddTorsor V P] diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 6343ce224c3da..fbf488bef44b3 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -344,7 +344,7 @@ theorem vectorSpan_isOrtho_altitude_direction {n : ℕ} (s : Simplex ℝ P (n + rw [direction_altitude] exact (Submodule.isOrtho_orthogonal_right _).mono_right inf_le_left -open FiniteDimensional +open Module /-- An altitude is finite-dimensional. -/ instance finiteDimensional_direction_altitude {n : ℕ} (s : Simplex ℝ P (n + 1)) (i : Fin (n + 2)) : @@ -392,7 +392,7 @@ theorem affineSpan_pair_eq_altitude_iff {n : ℕ} (s : Simplex ℝ P (n + 1)) (i rw [vectorSpan_eq_span_vsub_set_left_ne ℝ (Set.mem_insert _ _), Set.insert_diff_of_mem _ (Set.mem_singleton _), Set.diff_singleton_eq_self fun h => hne (Set.mem_singleton_iff.1 h), Set.image_singleton] - refine eq_of_le_of_finrank_eq ?_ ?_ + refine Submodule.eq_of_le_of_finrank_eq ?_ ?_ · rw [Submodule.span_le] simpa using h · rw [finrank_direction_altitude, finrank_span_set_eq_card] @@ -404,7 +404,7 @@ end Simplex namespace Triangle -open EuclideanGeometry Finset Simplex AffineSubspace FiniteDimensional +open EuclideanGeometry Finset Simplex AffineSubspace Module variable {V : Type*} {P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] [NormedAddTorsor V P] @@ -548,7 +548,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} have he : affineSpan ℝ (Set.range t₂.points) = affineSpan ℝ (Set.range t₁.points) := by refine ext_of_direction_eq ?_ ⟨t₁.points i₃, mem_affineSpan ℝ ⟨j₃, h₃⟩, mem_affineSpan ℝ (Set.mem_range_self _)⟩ - refine eq_of_le_of_finrank_eq (direction_le (spanPoints_subset_coe_of_subset_coe ?_)) ?_ + refine Submodule.eq_of_le_of_finrank_eq (direction_le (spanPoints_subset_coe_of_subset_coe ?_)) + ?_ · have hu : (Finset.univ : Finset (Fin 3)) = {j₁, j₂, j₃} := by clear h₁ h₂ h₃ -- Porting note (#11043): was `decide!` @@ -602,7 +603,7 @@ end Affine namespace EuclideanGeometry -open Affine AffineSubspace FiniteDimensional +open Affine AffineSubspace Module variable {V : Type*} {P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] [NormedAddTorsor V P] @@ -706,7 +707,7 @@ theorem affineSpan_of_orthocentricSystem {s : Set P} (ho : OrthocentricSystem s) ⟨p 0, mem_affineSpan ℝ (Set.mem_range_self _), mem_affineSpan ℝ (hps (Set.mem_range_self _))⟩ have hfd : FiniteDimensional ℝ (affineSpan ℝ s).direction := by rw [hs]; infer_instance haveI := hfd - refine eq_of_le_of_finrank_eq (direction_le (affineSpan_mono ℝ hps)) ?_ + refine Submodule.eq_of_le_of_finrank_eq (direction_le (affineSpan_mono ℝ hps)) ?_ rw [hs, direction_affineSpan, direction_affineSpan, ha.finrank_vectorSpan (Fintype.card_fin _), t.independent.finrank_vectorSpan (Fintype.card_fin _)] diff --git a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean index 20d5711b81e4d..6c5067aad4c23 100644 --- a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean @@ -33,7 +33,7 @@ namespace EuclideanGeometry variable {V : Type*} (P : Type*) -open FiniteDimensional +open Module /-- A `Sphere P` bundles a `center` and `radius`. This definition does not require the radius to be positive; that should be given as a hypothesis to lemmas that require it. -/ diff --git a/Mathlib/Geometry/Euclidean/Triangle.lean b/Mathlib/Geometry/Euclidean/Triangle.lean index 5a28ab76286ee..9e2ab6a89582d 100644 --- a/Mathlib/Geometry/Euclidean/Triangle.lean +++ b/Mathlib/Geometry/Euclidean/Triangle.lean @@ -295,7 +295,7 @@ theorem angle_add_angle_add_angle_eq_pi {p1 p2 p3 : P} (h2 : p2 ≠ p1) (h3 : p3 /-- The **sum of the angles of a triangle** (possibly degenerate, where the triangle is a line), oriented angles at point. -/ theorem oangle_add_oangle_add_oangle_eq_pi [Module.Oriented ℝ V (Fin 2)] - [Fact (FiniteDimensional.finrank ℝ V = 2)] {p1 p2 p3 : P} (h21 : p2 ≠ p1) (h32 : p3 ≠ p2) + [Fact (Module.finrank ℝ V = 2)] {p1 p2 p3 : P} (h21 : p2 ≠ p1) (h32 : p3 ≠ p2) (h13 : p1 ≠ p3) : ∡ p1 p2 p3 + ∡ p2 p3 p1 + ∡ p3 p1 p2 = π := by simpa only [neg_vsub_eq_vsub_rev] using positiveOrientation.oangle_add_cyc3_neg_left (vsub_ne_zero.mpr h21) (vsub_ne_zero.mpr h32) diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index d32ef29050ef2..16f7d52a36180 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -34,7 +34,7 @@ variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] -open Function Filter FiniteDimensional Set Metric +open Function Filter Module Set Metric open scoped Topology Manifold diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 6b82e8495e419..f5d59a6f08222 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -67,7 +67,7 @@ variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] noncomputable section -open Metric FiniteDimensional Function +open Metric Module Function open scoped Manifold diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 23d818e090332..07606636dad70 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -57,7 +57,7 @@ smooth bump function, partition of unity universe uι uE uH uM uF -open Function Filter FiniteDimensional Set +open Function Filter Module Set open scoped Topology Manifold noncomputable section diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index e99bb7ea3797f..3e3f2e4bb92fd 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -31,7 +31,7 @@ variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E [FiniteDimensional ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] -open Function Filter FiniteDimensional Set +open Function Filter Module Set open scoped Manifold noncomputable section diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index e0cc46543b743..935dc82abfb3b 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -29,14 +29,14 @@ section AffineSpace' variable (k : Type*) {V : Type*} {P : Type*} variable {ι : Type*} -open AffineSubspace FiniteDimensional Module +open AffineSubspace Module variable [DivisionRing k] [AddCommGroup V] [Module k V] [AffineSpace V P] /-- The `vectorSpan` of a finite set is finite-dimensional. -/ theorem finiteDimensional_vectorSpan_of_finite {s : Set P} (h : Set.Finite s) : FiniteDimensional k (vectorSpan k s) := - span_of_finite k <| h.vsub h + .span_of_finite k <| h.vsub h /-- The `vectorSpan` of a family indexed by a `Fintype` is finite-dimensional. -/ @@ -202,7 +202,7 @@ theorem finrank_vectorSpan_le_iff_not_affineIndependent [Fintype ι] (p : ι → variable {k} lemma AffineIndependent.card_le_finrank_succ [Fintype ι] {p : ι → P} (hp : AffineIndependent k p) : - Fintype.card ι ≤ FiniteDimensional.finrank k (vectorSpan k (Set.range p)) + 1 := by + Fintype.card ι ≤ Module.finrank k (vectorSpan k (Set.range p)) + 1 := by cases isEmpty_or_nonempty ι · simp [Fintype.card_eq_zero] rw [← tsub_le_iff_right] @@ -257,7 +257,7 @@ theorem AffineIndependent.vectorSpan_image_finset_eq_of_le_of_card_eq_finrank_ad (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) : vectorSpan k (s.image p : Set P) = sm := - eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan_image_finset hc + Submodule.eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan_image_finset hc /-- If the `vectorSpan` of a finite affinely independent family lies in a submodule with dimension one less than its @@ -266,7 +266,7 @@ theorem AffineIndependent.vectorSpan_eq_of_le_of_card_eq_finrank_add_one [Fintyp (hi : AffineIndependent k p) {sm : Submodule k V} [FiniteDimensional k sm] (hle : vectorSpan k (Set.range p) ≤ sm) (hc : Fintype.card ι = finrank k sm + 1) : vectorSpan k (Set.range p) = sm := - eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan hc + Submodule.eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan hc /-- If the `affineSpan` of a finite subset of an affinely independent family lies in an affine subspace whose direction has dimension one @@ -669,7 +669,7 @@ section DivisionRing variable {k : Type*} {V : Type*} {P : Type*} -open AffineSubspace FiniteDimensional Module +open AffineSubspace Module Module variable [DivisionRing k] [AddCommGroup V] [Module k V] [AffineSpace V P] @@ -764,12 +764,12 @@ protected theorem finite_set [FiniteDimensional k V] {s : Set ι} (b : AffineBas finite_set_of_fin_dim_affineIndependent k b.ind theorem card_eq_finrank_add_one [Fintype ι] (b : AffineBasis ι k P) : - Fintype.card ι = FiniteDimensional.finrank k V + 1 := + Fintype.card ι = Module.finrank k V + 1 := have : FiniteDimensional k V := b.finiteDimensional b.ind.affineSpan_eq_top_iff_card_eq_finrank_add_one.mp b.tot theorem exists_affineBasis_of_finiteDimensional [Fintype ι] [FiniteDimensional k V] - (h : Fintype.card ι = FiniteDimensional.finrank k V + 1) : Nonempty (AffineBasis ι k P) := by + (h : Fintype.card ι = Module.finrank k V + 1) : Nonempty (AffineBasis ι k P) := by obtain ⟨s, b, hb⟩ := AffineBasis.exists_affineBasis k V P lift s to Finset P using b.finite_set refine ⟨b.reindex <| Fintype.equivOfCardEq ?_⟩ diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index 169819f4632df..bff969f2c534c 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -302,7 +302,7 @@ lemma inf_orthogonal_self_le_ker_restrict {W : Submodule R M} (b₁ : B.IsRefl) variable [FiniteDimensional K V] -open FiniteDimensional Submodule +open Module Submodule variable {B : BilinForm K V} diff --git a/Mathlib/LinearAlgebra/Charpoly/Basic.lean b/Mathlib/LinearAlgebra/Charpoly/Basic.lean index eff887d7c175d..a54a424f7f68e 100644 --- a/Mathlib/LinearAlgebra/Charpoly/Basic.lean +++ b/Mathlib/LinearAlgebra/Charpoly/Basic.lean @@ -51,7 +51,7 @@ section Coeff theorem charpoly_monic : f.charpoly.Monic := Matrix.charpoly_monic _ -open FiniteDimensional in +open Module in lemma charpoly_natDegree [Nontrivial R] [StrongRankCondition R] : natDegree (charpoly f) = finrank R M := by rw [charpoly, Matrix.charpoly_natDegree_eq_dim, finrank_eq_card_chooseBasisIndex] diff --git a/Mathlib/LinearAlgebra/Coevaluation.lean b/Mathlib/LinearAlgebra/Coevaluation.lean index bd0df67099a43..31b64ef5ce5f8 100644 --- a/Mathlib/LinearAlgebra/Coevaluation.lean +++ b/Mathlib/LinearAlgebra/Coevaluation.lean @@ -25,7 +25,7 @@ noncomputable section section coevaluation -open TensorProduct FiniteDimensional +open TensorProduct Module open TensorProduct diff --git a/Mathlib/LinearAlgebra/Contraction.lean b/Mathlib/LinearAlgebra/Contraction.lean index 8ba1276a30496..0e7e5c811a68a 100644 --- a/Mathlib/LinearAlgebra/Contraction.lean +++ b/Mathlib/LinearAlgebra/Contraction.lean @@ -201,7 +201,7 @@ section CommRing variable [CommRing R] variable [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] [AddCommGroup Q] variable [Module R M] [Module R N] [Module R P] [Module R Q] -variable [Free R M] [Finite R M] [Free R N] [Finite R N] +variable [Free R M] [Module.Finite R M] [Free R N] [Module.Finite R N] /-- When `M` is a finite free module, the map `lTensorHomToHomLTensor` is an equivalence. Note that `lTensorHomEquivHomLTensor` is not defined directly in terms of diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index ce891c2dddd94..8f8edaca16e32 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -231,15 +231,15 @@ theorem det_id : LinearMap.det (LinearMap.id : M →ₗ[A] M) = 1 := @[simp] theorem det_smul {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] [Module 𝕜 M] (c : 𝕜) (f : M →ₗ[𝕜] M) : - LinearMap.det (c • f) = c ^ FiniteDimensional.finrank 𝕜 M * LinearMap.det f := by + LinearMap.det (c • f) = c ^ Module.finrank 𝕜 M * LinearMap.det f := by by_cases H : ∃ s : Finset M, Nonempty (Basis s 𝕜 M) · have : FiniteDimensional 𝕜 M := by rcases H with ⟨s, ⟨hs⟩⟩ exact FiniteDimensional.of_fintype_basis hs - simp only [← det_toMatrix (FiniteDimensional.finBasis 𝕜 M), LinearEquiv.map_smul, + simp only [← det_toMatrix (Module.finBasis 𝕜 M), LinearEquiv.map_smul, Fintype.card_fin, Matrix.det_smul] · classical - have : FiniteDimensional.finrank 𝕜 M = 0 := finrank_eq_zero_of_not_exists_basis H + have : Module.finrank 𝕜 M = 0 := finrank_eq_zero_of_not_exists_basis H simp [coe_det, H, this] theorem det_zero' {ι : Type*} [Finite ι] [Nonempty ι] (b : Basis ι A M) : @@ -253,7 +253,7 @@ and `0` otherwise. We give a formula that also works in infinite dimension, wher the determinant to be `1`. -/ @[simp] theorem det_zero {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] [Module 𝕜 M] : - LinearMap.det (0 : M →ₗ[𝕜] M) = (0 : 𝕜) ^ FiniteDimensional.finrank 𝕜 M := by + LinearMap.det (0 : M →ₗ[𝕜] M) = (0 : 𝕜) ^ Module.finrank 𝕜 M := by simp only [← zero_smul 𝕜 (1 : M →ₗ[𝕜] M), det_smul, mul_one, MonoidHom.map_one] theorem det_eq_one_of_subsingleton [Subsingleton M] (f : M →ₗ[R] M) : @@ -263,14 +263,14 @@ theorem det_eq_one_of_subsingleton [Subsingleton M] (f : M →ₗ[R] M) : exact Matrix.det_isEmpty theorem det_eq_one_of_finrank_eq_zero {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] - [Module 𝕜 M] (h : FiniteDimensional.finrank 𝕜 M = 0) (f : M →ₗ[𝕜] M) : + [Module 𝕜 M] (h : Module.finrank 𝕜 M = 0) (f : M →ₗ[𝕜] M) : LinearMap.det (f : M →ₗ[𝕜] M) = 1 := by classical refine @LinearMap.det_cases M _ 𝕜 _ _ _ (fun t => t = 1) f ?_ rfl intro s b have : IsEmpty s := by rw [← Fintype.card_eq_zero_iff] - exact (FiniteDimensional.finrank_eq_card_basis b).symm.trans h + exact (Module.finrank_eq_card_basis b).symm.trans h exact Matrix.det_isEmpty /-- Conjugating a linear map by a linear equiv does not change its determinant. -/ @@ -423,8 +423,8 @@ theorem LinearEquiv.coe_ofIsUnitDet {f : M →ₗ[R] M'} {v : Basis ι R M} {v' determinant is nonzero. -/ abbrev LinearMap.equivOfDetNeZero {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] [Module 𝕜 M] [FiniteDimensional 𝕜 M] (f : M →ₗ[𝕜] M) (hf : LinearMap.det f ≠ 0) : M ≃ₗ[𝕜] M := - have : IsUnit (LinearMap.toMatrix (FiniteDimensional.finBasis 𝕜 M) - (FiniteDimensional.finBasis 𝕜 M) f).det := by + have : IsUnit (LinearMap.toMatrix (Module.finBasis 𝕜 M) + (Module.finBasis 𝕜 M) f).det := by rw [LinearMap.det_toMatrix] exact isUnit_iff_ne_zero.2 hf LinearEquiv.ofIsUnitDet this diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 2a3358623db3a..9a7b997644973 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -35,7 +35,7 @@ universe u v v' u₁' w w' variable {R S : Type u} {M : Type v} {M' : Type v'} {M₁ : Type v} variable {ι : Type w} {ι' : Type w'} {η : Type u₁'} {φ : η → Type*} -open Cardinal Basis Submodule Function Set FiniteDimensional DirectSum +open Basis Cardinal DirectSum Function Module Set Submodule variable [Ring R] [CommRing S] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M₁] variable [Module R M] @@ -142,7 +142,7 @@ theorem rank_prod' : Module.rank R (M × M₁) = Module.rank R M + Module.rank R /-- The finrank of `M × M'` is `(finrank R M) + (finrank R M')`. -/ @[simp] -theorem FiniteDimensional.finrank_prod [Module.Finite R M] [Module.Finite R M'] : +theorem Module.finrank_prod [Module.Finite R M] [Module.Finite R M'] : finrank R (M × M') = finrank R M + finrank R M' := by simp [finrank, rank_lt_aleph0 R M, rank_lt_aleph0 R M'] @@ -209,7 +209,7 @@ theorem rank_matrix'' (m n : Type u) [Finite m] [Finite n] : open Fintype -namespace FiniteDimensional +namespace Module @[simp] theorem finrank_finsupp {ι : Type v} [Fintype ι] : finrank R (ι →₀ M) = card ι * finrank R M := by @@ -234,7 +234,7 @@ theorem finrank_directSum {ι : Type v} [Fintype ι] (M : ι → Type w) [∀ i theorem finrank_matrix (m n : Type*) [Fintype m] [Fintype n] : finrank R (Matrix m n R) = card m * card n := by simp [finrank] -end FiniteDimensional +end Module end Finsupp @@ -260,13 +260,13 @@ theorem rank_pi [Finite η] : Module.rank R (∀ i, φ i) = variable (R) /-- The finrank of `(ι → R)` is `Fintype.card ι`. -/ -theorem FiniteDimensional.finrank_pi {ι : Type v} [Fintype ι] : +theorem Module.finrank_pi {ι : Type v} [Fintype ι] : finrank R (ι → R) = Fintype.card ι := by simp [finrank] --TODO: this should follow from `LinearEquiv.finrank_eq`, that is over a field. /-- The finrank of a finite product is the sum of the finranks. -/ -theorem FiniteDimensional.finrank_pi_fintype +theorem Module.finrank_pi_fintype {ι : Type v} [Fintype ι] {M : ι → Type w} [∀ i : ι, AddCommGroup (M i)] [∀ i : ι, Module R (M i)] [∀ i : ι, Module.Free R (M i)] [∀ i : ι, Module.Finite R (M i)] : finrank R (∀ i, M i) = ∑ i, finrank R (M i) := by @@ -294,12 +294,12 @@ variable (R) /-- The vector space of functions on a `Fintype ι` has finrank equal to the cardinality of `ι`. -/ @[simp] -theorem FiniteDimensional.finrank_fintype_fun_eq_card : finrank R (η → R) = Fintype.card η := +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 FiniteDimensional.finrank_fin_fun {n : ℕ} : finrank R (Fin n → R) = n := by simp +theorem Module.finrank_fin_fun {n : ℕ} : finrank R (Fin n → R) = n := by simp variable {R} @@ -343,7 +343,7 @@ theorem rank_tensorProduct' : /-- The `S`-finrank of `M ⊗[R] M'` is `(finrank S M) * (finrank R M')`. -/ @[simp] -theorem FiniteDimensional.finrank_tensorProduct : +theorem Module.finrank_tensorProduct : finrank R (M ⊗[S] M') = finrank R M * finrank S M' := by simp [finrank] end TensorProduct @@ -352,7 +352,7 @@ section SubmoduleRank section -open FiniteDimensional +open Module namespace Submodule @@ -413,7 +413,7 @@ theorem rank_span_finset_le (s : Finset M) : Module.rank R (span R (s : Set M)) theorem rank_span_of_finset (s : Finset M) : Module.rank R (span R (s : Set M)) < ℵ₀ := (rank_span_finset_le s).trans_lt (Cardinal.nat_lt_aleph0 _) -open Submodule FiniteDimensional +open Submodule Module variable (R) diff --git a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean index 81356f0bed052..cf1afda61d172 100644 --- a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean +++ b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean @@ -12,7 +12,7 @@ import Mathlib.LinearAlgebra.Dimension.RankNullity /-! # Dimension of vector spaces -In this file we provide results about `Module.rank` and `FiniteDimensional.finrank` of vector spaces +In this file we provide results about `Module.rank` and `Module.finrank` of vector spaces over division rings. ## Main statements @@ -112,7 +112,7 @@ end Module section Basis -open FiniteDimensional +open Module variable [DivisionRing K] [AddCommGroup V] [Module K V] diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 32cbe00cac733..5e70194888971 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -26,7 +26,7 @@ variable [Module R M] [Module R M'] [Module R M₁] attribute [local instance] nontrivial_of_invariantBasisNumber -open Cardinal Basis Submodule Function Set FiniteDimensional +open Basis Cardinal Function Module Set Submodule theorem rank_le {n : ℕ} (H : ∀ s : Finset M, (LinearIndependent R fun i : s => (i : M)) → s.card ≤ n) : @@ -363,7 +363,7 @@ variable [Nontrivial R] /-- A (finite dimensional) space that is a subsingleton has zero `finrank`. -/ @[nontriviality] -theorem FiniteDimensional.finrank_zero_of_subsingleton [Subsingleton M] : +theorem Module.finrank_zero_of_subsingleton [Subsingleton M] : finrank R M = 0 := by rw [finrank, rank_subsingleton', _root_.map_zero] @@ -374,12 +374,12 @@ section variable [NoZeroSMulDivisors R M] /-- A finite dimensional space is nontrivial if it has positive `finrank`. -/ -theorem FiniteDimensional.nontrivial_of_finrank_pos (h : 0 < finrank R M) : Nontrivial M := +theorem Module.nontrivial_of_finrank_pos (h : 0 < finrank R M) : Nontrivial M := rank_pos_iff_nontrivial.mp (lt_rank_of_lt_finrank h) /-- A finite dimensional space is nontrivial if it has `finrank` equal to the successor of a natural number. -/ -theorem FiniteDimensional.nontrivial_of_finrank_eq_succ {n : ℕ} +theorem Module.nontrivial_of_finrank_eq_succ {n : ℕ} (hn : finrank R M = n.succ) : Nontrivial M := nontrivial_of_finrank_pos (R := R) (by rw [hn]; exact n.succ_pos) @@ -398,31 +398,31 @@ section StrongRankCondition variable [StrongRankCondition R] [Module.Finite R M] /-- A finite rank torsion-free module has positive `finrank` iff it has a nonzero element. -/ -theorem FiniteDimensional.finrank_pos_iff_exists_ne_zero [NoZeroSMulDivisors R M] : +theorem Module.finrank_pos_iff_exists_ne_zero [NoZeroSMulDivisors R M] : 0 < finrank R M ↔ ∃ x : M, x ≠ 0 := by rw [← @rank_pos_iff_exists_ne_zero R M, ← finrank_eq_rank] norm_cast /-- An `R`-finite torsion-free module has positive `finrank` iff it is nontrivial. -/ -theorem FiniteDimensional.finrank_pos_iff [NoZeroSMulDivisors R M] : +theorem Module.finrank_pos_iff [NoZeroSMulDivisors R M] : 0 < finrank R M ↔ Nontrivial M := by rw [← rank_pos_iff_nontrivial (R := R), ← finrank_eq_rank] norm_cast /-- A nontrivial finite dimensional space has positive `finrank`. -/ -theorem FiniteDimensional.finrank_pos [NoZeroSMulDivisors R M] [h : Nontrivial M] : +theorem Module.finrank_pos [NoZeroSMulDivisors R M] [h : Nontrivial M] : 0 < finrank R M := finrank_pos_iff.mpr h -/-- See `FiniteDimensional.finrank_zero_iff` +/-- See `Module.finrank_zero_iff` for the stronger version with `NoZeroSMulDivisors R M`. -/ -theorem FiniteDimensional.finrank_eq_zero_iff : +theorem Module.finrank_eq_zero_iff : finrank R M = 0 ↔ ∀ x : M, ∃ a : R, a ≠ 0 ∧ a • x = 0 := by rw [← rank_eq_zero_iff (R := R), ← finrank_eq_rank] norm_cast /-- The `StrongRankCondition` is automatic. See `commRing_strongRankCondition`. -/ -theorem FiniteDimensional.finrank_eq_zero_iff_isTorsion {R} [CommRing R] [StrongRankCondition R] +theorem Module.finrank_eq_zero_iff_isTorsion {R} [CommRing R] [StrongRankCondition R] [IsDomain R] [Module R M] [Module.Finite R M] : finrank R M = 0 ↔ Module.IsTorsion R M := by rw [← rank_eq_zero_iff_isTorsion (R := R), ← finrank_eq_rank] @@ -430,14 +430,14 @@ theorem FiniteDimensional.finrank_eq_zero_iff_isTorsion {R} [CommRing R] [Strong /-- A finite dimensional space has zero `finrank` iff it is a subsingleton. This is the `finrank` version of `rank_zero_iff`. -/ -theorem FiniteDimensional.finrank_zero_iff [NoZeroSMulDivisors R M] : +theorem Module.finrank_zero_iff [NoZeroSMulDivisors R M] : finrank R M = 0 ↔ Subsingleton M := by rw [← rank_zero_iff (R := R), ← finrank_eq_rank] norm_cast end StrongRankCondition -theorem FiniteDimensional.finrank_eq_zero_of_rank_eq_zero (h : Module.rank R M = 0) : +theorem Module.finrank_eq_zero_of_rank_eq_zero (h : Module.rank R M = 0) : finrank R M = 0 := by delta finrank rw [h, zero_toNat] diff --git a/Mathlib/LinearAlgebra/Dimension/Finrank.lean b/Mathlib/LinearAlgebra/Dimension/Finrank.lean index 9bf3954c81775..69668ef8b667a 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finrank.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finrank.lean @@ -13,7 +13,7 @@ Definition of the rank of a module, or dimension of a vector space, as a natural ## Main definitions -Defined is `FiniteDimensional.finrank`, the dimension of a finite dimensional space, returning a +Defined is `Module.finrank`, the dimension of a finite dimensional space, returning a `Nat`, as opposed to `Module.rank`, which returns a `Cardinal`. When the space has infinite dimension, its `finrank` is by convention set to `0`. @@ -38,7 +38,7 @@ open Cardinal Submodule Module Function variable {R : Type u} {M : Type v} {N : Type w} variable [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -namespace FiniteDimensional +namespace Module section Ring @@ -52,6 +52,8 @@ of `M` over `R`. noncomputable def finrank (R M : Type*) [Semiring R] [AddCommGroup M] [Module R M] : ℕ := Cardinal.toNat (Module.rank R M) +@[deprecated (since := "2024-10-01")] protected alias _root_.FiniteDimensional.finrank := finrank + theorem finrank_eq_of_rank_eq {n : ℕ} (h : Module.rank R M = ↑n) : finrank R M = n := by apply_fun toNat at h rw [toNat_natCast] at h @@ -92,9 +94,9 @@ theorem finrank_le_finrank_of_rank_le_rank end Ring -end FiniteDimensional +end Module -open FiniteDimensional +open Module namespace LinearEquiv diff --git a/Mathlib/LinearAlgebra/Dimension/Free.lean b/Mathlib/LinearAlgebra/Dimension/Free.lean index 3b06f6f965e75..af3007b4da717 100644 --- a/Mathlib/LinearAlgebra/Dimension/Free.lean +++ b/Mathlib/LinearAlgebra/Dimension/Free.lean @@ -14,7 +14,7 @@ import Mathlib.SetTheory.Cardinal.Finsupp ## Main result - `LinearEquiv.nonempty_equiv_iff_lift_rank_eq`: Two free modules are isomorphic iff they have the same dimension. -- `FiniteDimensional.finBasis`: +- `Module.finBasis`: An arbitrary basis of a finite free module indexed by `Fin n` given `finrank R M = n`. -/ @@ -24,7 +24,7 @@ noncomputable section universe u v v' w -open Cardinal Basis Submodule Function Set DirectSum FiniteDimensional +open Cardinal Basis Submodule Function Set DirectSum Module section Tower @@ -57,7 +57,7 @@ theorem rank_mul_rank (A : Type v) [AddCommGroup A] /-- Tower law: if `A` is a `K`-module and `K` is an extension of `F` then $\operatorname{rank}_F(A) = \operatorname{rank}_F(K) * \operatorname{rank}_K(A)$. -/ -theorem FiniteDimensional.finrank_mul_finrank : finrank F K * finrank K A = finrank F A := by +theorem Module.finrank_mul_finrank : finrank F K * finrank K A = finrank F A := by simp_rw [finrank] rw [← toNat_lift.{w} (Module.rank F K), ← toNat_lift.{v} (Module.rank K A), ← toNat_mul, lift_rank_mul_lift_rank, toNat_lift] @@ -79,7 +79,7 @@ theorem rank_eq_card_chooseBasisIndex : Module.rank R M = #(ChooseBasisIndex R M (chooseBasis R M).mk_eq_rank''.symm /-- The finrank of a free module `M` over `R` is the cardinality of `ChooseBasisIndex R M`. -/ -theorem _root_.FiniteDimensional.finrank_eq_card_chooseBasisIndex [Module.Finite R M] : +theorem _root_.Module.finrank_eq_card_chooseBasisIndex [Module.Finite R M] : finrank R M = Fintype.card (ChooseBasisIndex R M) := by simp [finrank, rank_eq_card_chooseBasisIndex] @@ -161,35 +161,30 @@ noncomputable def LinearEquiv.ofFinrankEq [Module.Finite R M] [Module.Finite R M variable {M M'} +namespace Module + /-- See `rank_lt_aleph0` for the inverse direction without `Module.Free R M`. -/ -lemma Module.rank_lt_alpeh0_iff : - Module.rank R M < ℵ₀ ↔ Module.Finite R M := by +lemma rank_lt_aleph0_iff : Module.rank R M < ℵ₀ ↔ Module.Finite R M := by rw [Free.rank_eq_card_chooseBasisIndex, mk_lt_aleph0_iff] exact ⟨fun h ↦ Finite.of_basis (Free.chooseBasis R M), fun I ↦ Finite.of_fintype (Free.ChooseBasisIndex R M)⟩ -theorem FiniteDimensional.finrank_of_not_finite - (h : ¬Module.Finite R M) : - finrank R M = 0 := by - rw [finrank, toNat_eq_zero, ← not_lt, Module.rank_lt_alpeh0_iff] +theorem finrank_of_not_finite (h : ¬Module.Finite R M) : finrank R M = 0 := by + rw [finrank, toNat_eq_zero, ← not_lt, Module.rank_lt_aleph0_iff] exact .inr h -theorem Module.finite_of_finrank_pos (h : 0 < finrank R M) : - Module.Finite R M := by +theorem finite_of_finrank_pos (h : 0 < finrank R M) : Module.Finite R M := by contrapose h simp [finrank_of_not_finite h] -theorem Module.finite_of_finrank_eq_succ {n : ℕ} - (hn : finrank R M = n.succ) : Module.Finite R M := - Module.finite_of_finrank_pos <| by rw [hn]; exact n.succ_pos +theorem finite_of_finrank_eq_succ {n : ℕ} (hn : finrank R M = n.succ) : Module.Finite R M := + finite_of_finrank_pos <| by rw [hn]; exact n.succ_pos -theorem Module.finite_iff_of_rank_eq_nsmul {W} [AddCommGroup W] - [Module R W] [Module.Free R W] {n : ℕ} (hn : n ≠ 0) - (hVW : Module.rank R M = n • Module.rank R W) : +theorem finite_iff_of_rank_eq_nsmul {W} [AddCommGroup W] [Module R W] [Module.Free R W] {n : ℕ} + (hn : n ≠ 0) (hVW : Module.rank R M = n • Module.rank R W) : Module.Finite R M ↔ Module.Finite R W := by - simp only [← rank_lt_alpeh0_iff, hVW, nsmul_lt_aleph0_iff_of_ne_zero hn] + simp only [← rank_lt_aleph0_iff, hVW, nsmul_lt_aleph0_iff_of_ne_zero hn] -namespace FiniteDimensional variable (R M) /-- A finite rank free module has a basis indexed by `Fin (finrank R M)`. -/ @@ -220,4 +215,4 @@ theorem basisUnique_repr_eq_zero_iff {ι : Type*} [Unique ι] (basisUnique ι h).repr.map_eq_zero_iff.mp (Finsupp.ext fun j => Subsingleton.elim i j ▸ hv), fun hv => by rw [hv, LinearEquiv.map_zero, Finsupp.zero_apply]⟩ -end FiniteDimensional +end Module diff --git a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean index 507f9e3b737de..17254b66084e3 100644 --- a/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/FreeAndStrongRankCondition.lean @@ -17,7 +17,7 @@ and `Mathlib/LinearAlgebra/FiniteDimensional.lean`. -/ -open Cardinal Submodule Set FiniteDimensional +open Cardinal Module Module Set Submodule universe u v @@ -27,7 +27,7 @@ variable {K : Type u} {V : Type v} [Ring K] [StrongRankCondition K] [AddCommGrou /-- The `ι` indexed basis on `V`, where `ι` is an empty type and `V` is zero-dimensional. -See also `FiniteDimensional.finBasis`. +See also `Module.finBasis`. -/ noncomputable def Basis.ofRankEqZero [Module.Free K V] {ι : Type*} [IsEmpty ι] (hV : Module.rank K V = 0) : Basis ι K V := @@ -186,7 +186,7 @@ theorem finrank_eq_one_iff [Module.Free K V] (ι : Type*) [Unique ι] : finrank K V = 1 ↔ Nonempty (Basis ι K V) := by constructor · intro h - exact ⟨basisUnique ι h⟩ + exact ⟨Module.basisUnique ι h⟩ · rintro ⟨b⟩ simpa using finrank_eq_card_basis b diff --git a/Mathlib/LinearAlgebra/Dimension/Localization.lean b/Mathlib/LinearAlgebra/Dimension/Localization.lean index 215eca3c7a657..b6445fadfe529 100644 --- a/Mathlib/LinearAlgebra/Dimension/Localization.lean +++ b/Mathlib/LinearAlgebra/Dimension/Localization.lean @@ -15,9 +15,9 @@ import Mathlib.RingTheory.OreLocalization.OreSet - `IsLocalizedModule.lift_rank_eq`: `rank_Rₚ Mₚ = rank R M`. - `rank_quotient_add_rank_of_isDomain`: The **rank-nullity theorem** for commutative domains. - -/ -open Cardinal nonZeroDivisors + +open Cardinal Module nonZeroDivisors section CommRing diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index 1e94b3ea6e179..2e6cb6ba25096 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -171,7 +171,7 @@ theorem Submodule.rank_add_le_rank_add_rank (s t : Submodule R M) : section Finrank -open Submodule FiniteDimensional +open Submodule Module variable [StrongRankCondition R] diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index aecd0ad209da8..12001f79238c5 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -376,14 +376,12 @@ theorem Ideal.rank_eq {R S : Type*} [CommRing R] [StrongRankCondition R] [Ring S ((LinearMap.ker_eq_bot (f := (Submodule.subtype I : I →ₗ[R] S))).mpr Subtype.coe_injective))) (c.card_le_card_of_linearIndependent this) -open FiniteDimensional +namespace Module theorem finrank_eq_nat_card_basis (h : Basis ι R M) : finrank R M = Nat.card ι := by rw [Nat.card, ← toNat_lift.{v}, h.mk_eq_rank, toNat_lift, finrank] -namespace FiniteDimensional - /-- If a vector space (or module) has a finite basis, then its dimension (or rank) is equal to the cardinality of the basis. -/ theorem finrank_eq_card_basis {ι : Type w} [Fintype ι] (h : Basis ι R M) : @@ -392,8 +390,8 @@ theorem finrank_eq_card_basis {ι : Type w} [Fintype ι] (h : Basis ι R M) : /-- If a free module is of finite rank, then the cardinality of any basis is equal to its `finrank`. -/ -theorem _root_.Module.mk_finrank_eq_card_basis [Module.Finite R M] - {ι : Type w} (h : Basis ι R M) : (finrank R M : Cardinal.{w}) = #ι := by +theorem mk_finrank_eq_card_basis [Module.Finite R M] {ι : Type w} (h : Basis ι R M) : + (finrank R M : Cardinal.{w}) = #ι := by cases @nonempty_fintype _ (Module.Finite.finite_basis h) rw [Cardinal.mk_fintype, finrank_eq_card_basis h] @@ -402,10 +400,6 @@ cardinality of the basis. This lemma uses a `Finset` instead of indexed types. - theorem finrank_eq_card_finset_basis {ι : Type w} {b : Finset ι} (h : Basis b R M) : finrank R M = Finset.card b := by rw [finrank_eq_card_basis h, Fintype.card_coe] -end FiniteDimensional - -open FiniteDimensional - variable (R) @[simp] @@ -415,15 +409,15 @@ theorem rank_self : Module.rank R R = 1 := by /-- A ring satisfying `StrongRankCondition` (such as a `DivisionRing`) is one-dimensional as a module over itself. -/ @[simp] -theorem FiniteDimensional.finrank_self : finrank R R = 1 := +theorem finrank_self : finrank R R = 1 := finrank_eq_of_rank_eq (by simp) /-- Given a basis of a ring over itself indexed by a type `ι`, then `ι` is `Unique`. -/ -noncomputable def Basis.unique {ι : Type*} (b : Basis ι R R) : Unique ι := by - have A : Cardinal.mk ι = ↑(FiniteDimensional.finrank R R) := +noncomputable def _root_.Basis.unique {ι : Type*} (b : Basis ι R R) : Unique ι := by + have A : Cardinal.mk ι = ↑(Module.finrank R R) := (Module.mk_finrank_eq_card_basis b).symm -- Porting note: replace `algebraMap.coe_one` with `Nat.cast_one` - simp only [Cardinal.eq_one_iff_unique, FiniteDimensional.finrank_self, Nat.cast_one] at A + simp only [Cardinal.eq_one_iff_unique, Module.finrank_self, Nat.cast_one] at A exact Nonempty.some ((unique_iff_subsingleton_and_nonempty _).2 A) variable (M) @@ -436,19 +430,15 @@ theorem rank_lt_aleph0 [Module.Finite R M] : Module.rank R M < ℵ₀ := by refine (ciSup_le' fun i => ?_).trans_lt (nat_lt_aleph0 S.card) exact linearIndependent_le_span_finset _ i.prop S hS -@[deprecated (since := "2024-01-01")] -protected alias FiniteDimensional.rank_lt_aleph0 := rank_lt_aleph0 - /-- If `M` is finite, `finrank M = rank M`. -/ @[simp] -theorem finrank_eq_rank [Module.Finite R M] : - ↑(FiniteDimensional.finrank R M) = Module.rank R M := by - rw [FiniteDimensional.finrank, cast_toNat_of_lt_aleph0 (rank_lt_aleph0 R M)] +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)] + +end Module -@[deprecated (since := "2024-01-01")] -protected alias FiniteDimensional.finrank_eq_rank := finrank_eq_rank +open Module -variable {R M} variable {M'} [AddCommGroup M'] [Module R M'] theorem LinearMap.finrank_le_finrank_of_injective [Module.Finite R M'] {f : M →ₗ[R] M'} diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index b4d2e69e8d184..1d2a2669cab75 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -91,6 +91,8 @@ The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to splitting of `V₁`. -/ +open Module Submodule + noncomputable section namespace Module @@ -384,7 +386,7 @@ theorem toDualEquiv_apply (m : M) : b.toDualEquiv m = b.toDual m := theorem linearEquiv_dual_iff_finiteDimensional [Field K] [AddCommGroup V] [Module K V] : Nonempty (V ≃ₗ[K] Dual K V) ↔ FiniteDimensional K V := by refine ⟨fun ⟨e⟩ ↦ ?_, fun h ↦ ⟨(Module.Free.chooseBasis K V).toDualEquiv⟩⟩ - rw [FiniteDimensional, ← Module.rank_lt_alpeh0_iff] + rw [FiniteDimensional, ← Module.rank_lt_aleph0_iff] by_contra! apply (lift_rank_lt_rank_dual this).ne have := e.lift_rank_eq @@ -449,12 +451,12 @@ theorem eval_range {ι : Type*} [Finite ι] (b : Basis ι R M) : section -variable [Finite R M] [Free R M] +variable [Module.Finite R M] [Free R M] instance dual_free : Free R (Dual R M) := Free.of_basis (Free.chooseBasis R M).dualBasis -instance dual_finite : Finite R (Dual R M) := +instance dual_finite : Module.Finite R (Dual R M) := Finite.of_basis (Free.chooseBasis R M).dualBasis end @@ -482,7 +484,7 @@ universe uK uV variable {K : Type uK} {V : Type uV} variable [CommRing K] [AddCommGroup V] [Module K V] [Module.Free K V] -open Module Module.Dual Submodule LinearMap Cardinal Basis FiniteDimensional +open Module Module.Dual Submodule LinearMap Cardinal Basis Module section @@ -540,7 +542,7 @@ theorem nontrivial_dual_iff : instance instNontrivialDual [Nontrivial V] : Nontrivial (Dual K V) := (nontrivial_dual_iff K).mpr inferInstance -theorem finite_dual_iff : Finite K (Dual K V) ↔ Finite K V := by +theorem finite_dual_iff : Module.Finite K (Dual K V) ↔ Module.Finite K V := by constructor <;> intro h · obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) nontriviality K @@ -578,7 +580,7 @@ class IsReflexive : Prop where lemma bijective_dual_eval [IsReflexive R M] : Bijective (Dual.eval R M) := IsReflexive.bijective_dual_eval' -instance IsReflexive.of_finite_of_free [Finite R M] [Free R M] : IsReflexive R M where +instance IsReflexive.of_finite_of_free [Module.Finite R M] [Free R M] : IsReflexive R M where bijective_dual_eval' := ⟨LinearMap.ker_eq_bot.mp (Free.chooseBasis R M).eval_ker, LinearMap.range_eq_top.mp (Free.chooseBasis R M).eval_range⟩ @@ -1057,7 +1059,7 @@ theorem dualEquivDual_apply (φ : Module.Dual K W) : section -open FiniteDimensional +open FiniteDimensional Module instance instModuleDualFiniteDimensional [FiniteDimensional K V] : FiniteDimensional K (Module.Dual K V) := by @@ -1096,7 +1098,7 @@ noncomputable def quotEquivAnnihilator (W : Subspace K V) : (V ⧸ W) ≃ₗ[K] -- refine' LinearEquiv.quot_equiv_of_equiv _ (Basis.ofVectorSpace K V).toDualEquiv -- exact (Basis.ofVectorSpace K W).toDualEquiv.trans W.dual_equiv_dual -open FiniteDimensional +open Module @[simp] theorem finrank_dualCoannihilator_eq {Φ : Subspace K (Module.Dual K V)} : @@ -1231,7 +1233,7 @@ theorem dualQuotEquivDualAnnihilator_symm_apply_mk (W : Submodule R M) (φ : W.d rfl theorem finite_dualAnnihilator_iff {W : Submodule R M} [Free R (M ⧸ W)] : - Finite R W.dualAnnihilator ↔ Finite R (M ⧸ W) := + Module.Finite R W.dualAnnihilator ↔ Module.Finite R (M ⧸ W) := (Finite.equiv_iff W.dualQuotEquivDualAnnihilator.symm).trans (finite_dual_iff R) open LinearMap in @@ -1324,7 +1326,6 @@ lemma range_eq_top_of_ne_zero : rw [eq_top_iff] exact fun x _ ↦ ⟨x • (f v)⁻¹ • v, by simp [inv_mul_cancel₀ hv]⟩ -open FiniteDimensional variable [FiniteDimensional K V₁] lemma finrank_ker_add_one_of_ne_zero : @@ -1479,7 +1480,7 @@ end Subspace section FiniteDimensional -open FiniteDimensional LinearMap +open Module LinearMap namespace LinearMap @@ -1610,7 +1611,7 @@ theorem dualAnnihilator_dualAnnihilator_eq_map (W : Subspace K V) [FiniteDimensi haveI := e1.finiteDimensional let e2 := (Free.chooseBasis K _).toDualEquiv ≪≫ₗ W.dualAnnihilator.dualQuotEquivDualAnnihilator haveI := LinearEquiv.finiteDimensional (V₂ := W.dualAnnihilator.dualAnnihilator) e2 - rw [FiniteDimensional.eq_of_le_of_finrank_eq (map_le_dualAnnihilator_dualAnnihilator W)] + rw [eq_of_le_of_finrank_eq (map_le_dualAnnihilator_dualAnnihilator W)] rw [← (equivMapOfInjective _ (eval_apply_injective K (V := V)) W).finrank_eq, e1.finrank_eq] exact e2.finrank_eq diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 1495ca7c3ce48..3effed50341ef 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -53,7 +53,7 @@ namespace Module namespace End -open FiniteDimensional Set +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] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean index 561d462974d51..e6cf66f99d9ea 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Minpoly.lean @@ -21,7 +21,7 @@ namespace Module namespace End -open Polynomial FiniteDimensional +open Polynomial Module open scoped Polynomial diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index a6fb4ee72e490..c0f53eccf53f8 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -38,7 +38,7 @@ generalized eigenspaces span the whole space. eigenspace, eigenvector, eigenvalue, eigen -/ -open Set Function Module FiniteDimensional +open Set Function Module Module variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index f2439bd876a5f..30837353e1b0a 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -33,7 +33,7 @@ variable {R K M : Type*} [CommRing R] [IsDomain R] [Field K] [AddCommGroup M] variable [Module R M] [Module.Finite R M] [Module.Free R M] variable [Module K M] [Module.Finite K M] -open FiniteDimensional Module.Free Polynomial +open Module Module.Free Polynomial lemma IsNilpotent.charpoly_eq_X_pow_finrank (φ : Module.End R M) (h : IsNilpotent φ) : φ.charpoly = X ^ finrank R M := by diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 8aec07c6ff9dd..7b2f68f6cdcdd 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -27,7 +27,7 @@ variable {K : Type u} {V : Type v} namespace Submodule -open IsNoetherian FiniteDimensional +open IsNoetherian Module section DivisionRing @@ -116,7 +116,7 @@ end FiniteDimensional namespace LinearMap -open FiniteDimensional +open Module section DivisionRing @@ -142,7 +142,7 @@ end DivisionRing end LinearMap -open FiniteDimensional +open Module namespace LinearMap diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 7802cfc81de05..acd30d6e430a9 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -27,7 +27,7 @@ that all these points of view are equivalent, with the following lemmas - `fintypeBasisIndex` states that a finite-dimensional vector space has a finite basis -- `FiniteDimensional.finBasis` and `FiniteDimensional.finBasisOfFinrankEq` +- `Module.finBasis` and `Module.finBasisOfFinrankEq` are bases for finite dimensional vector spaces, where the index type is `Fin` (in `Mathlib.LinearAlgebra.Dimension.Free`) - `of_fintype_basis` states that the existence of a basis indexed by a @@ -70,7 +70,7 @@ Plenty of the results hold for general fg modules or notherian modules, and they universe u v v' w -open Cardinal Submodule Module Function +open Cardinal Function IsNoetherian Module Submodule /-- `FiniteDimensional` vector spaces are defined to be finite modules. Use `FiniteDimensional.of_fintype_basis` to prove finite dimension from another definition. -/ @@ -80,11 +80,6 @@ abbrev FiniteDimensional (K V : Type*) [DivisionRing K] [AddCommGroup V] [Module variable {K : Type u} {V : Type v} namespace FiniteDimensional - -open IsNoetherian - -section DivisionRing - variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] [Module K V₂] @@ -145,10 +140,8 @@ theorem of_finite_basis {ι : Type w} {s : Set ι} (h : Basis s K V) (hs : Set.F instance finiteDimensional_submodule [FiniteDimensional K V] (S : Submodule K V) : FiniteDimensional K S := by letI : IsNoetherian K V := iff_fg.2 ?_ - · exact - iff_fg.1 - (IsNoetherian.iff_rank_lt_aleph0.2 - ((Submodule.rank_le _).trans_lt (_root_.rank_lt_aleph0 K V))) + · exact iff_fg.1 <| IsNoetherian.iff_rank_lt_aleph0.2 <| + (Submodule.rank_le _).trans_lt (rank_lt_aleph0 K V) · infer_instance /-- A quotient of a finite-dimensional space is also finite-dimensional. -/ @@ -156,18 +149,6 @@ instance finiteDimensional_quotient [FiniteDimensional K V] (S : Submodule K V) FiniteDimensional K (V ⧸ S) := Module.Finite.quotient K S -variable (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. -/ -theorem finrank_eq_rank' [FiniteDimensional K V] : (finrank K V : Cardinal.{v}) = Module.rank K V := - finrank_eq_rank _ _ - -variable {K V} - -theorem finrank_of_infinite_dimensional (h : ¬FiniteDimensional K V) : finrank K V = 0 := - FiniteDimensional.finrank_of_not_finite h - theorem of_finrank_pos (h : 0 < finrank K V) : FiniteDimensional K V := Module.finite_of_finrank_pos h @@ -181,6 +162,24 @@ theorem of_fact_finrank_eq_succ (n : ℕ) [hn : Fact (finrank K V = n + 1)] : FiniteDimensional K V := of_finrank_eq_succ hn.out +end FiniteDimensional + +namespace Module + +variable (K V) +variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [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. -/ +theorem finrank_eq_rank' [FiniteDimensional K V] : (finrank K V : Cardinal.{v}) = Module.rank K V := + finrank_eq_rank _ _ + +variable {K V} + +theorem finrank_of_infinite_dimensional (h : ¬FiniteDimensional K V) : finrank K V = 0 := + Module.finrank_of_not_finite h + theorem finiteDimensional_iff_of_rank_eq_nsmul {W} [AddCommGroup W] [Module K W] {n : ℕ} (hn : n ≠ 0) (hVW : Module.rank K V = n • Module.rank K W) : FiniteDimensional K V ↔ FiniteDimensional K W := @@ -192,6 +191,13 @@ theorem finrank_eq_card_basis' [FiniteDimensional K V] {ι : Type w} (h : Basis (finrank K V : Cardinal.{w}) = #ι := Module.mk_finrank_eq_card_basis h +end Module + +namespace FiniteDimensional +section DivisionRing +variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] + [Module K V₂] + theorem _root_.LinearIndependent.lt_aleph0_of_finiteDimensional {ι : Type w} [FiniteDimensional K V] {v : ι → V} (h : LinearIndependent K v) : #ι < ℵ₀ := h.lt_aleph0_of_finite @@ -254,8 +260,6 @@ section open Finset -section - variable {L : Type*} [LinearOrderedField L] variable {W : Type v} [AddCommGroup W] [Module L W] @@ -271,16 +275,14 @@ theorem exists_relation_sum_zero_pos_coefficient_of_finrank_succ_lt_card [Finite exact ⟨f, sum, total, exists_pos_of_sum_zero_of_exists_nonzero f total nonzero⟩ -end - end /-- In a vector space with dimension 1, each set {v} is a basis for `v ≠ 0`. -/ @[simps repr_apply] noncomputable def basisSingleton (ι : Type*) [Unique ι] (h : finrank K V = 1) (v : V) (hv : v ≠ 0) : Basis ι K V := - let b := FiniteDimensional.basisUnique ι h - let h : b.repr v default ≠ 0 := mt FiniteDimensional.basisUnique_repr_eq_zero_iff.mp hv + let b := Module.basisUnique ι h + let h : b.repr v default ≠ 0 := mt Module.basisUnique_repr_eq_zero_iff.mp hv Basis.ofRepr { toFun := fun w => Finsupp.single default (b.repr w default / b.repr v default) invFun := fun f => f default • v @@ -326,8 +328,6 @@ section ZeroRank variable [DivisionRing K] [AddCommGroup V] [Module K V] -open FiniteDimensional - theorem FiniteDimensional.of_rank_eq_nat {n : ℕ} (h : Module.rank K V = n) : FiniteDimensional K V := Module.finite_of_rank_eq_nat h @@ -350,7 +350,7 @@ alias finiteDimensional_of_rank_eq_one := FiniteDimensional.of_rank_eq_one variable (K V) instance finiteDimensional_bot : FiniteDimensional K (⊥ : Submodule K V) := - of_rank_eq_zero <| by simp + .of_rank_eq_zero <| by simp variable {K V} @@ -358,7 +358,7 @@ end ZeroRank namespace Submodule -open IsNoetherian FiniteDimensional +open IsNoetherian Module section DivisionRing @@ -423,7 +423,7 @@ end Submodule namespace LinearEquiv -open FiniteDimensional +open Module variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] [Module K V₂] @@ -448,10 +448,7 @@ instance finiteDimensional_finsupp {ι : Type*} [Finite ι] [FiniteDimensional K end -namespace FiniteDimensional - -section DivisionRing - +namespace Submodule variable [DivisionRing K] [AddCommGroup V] [Module K V] {V₂ : Type v'} [AddCommGroup V₂] [Module K V₂] @@ -469,33 +466,29 @@ theorem eq_of_le_of_finrank_eq {S₁ S₂ : Submodule K V} [FiniteDimensional K (hd : finrank K S₁ = finrank K S₂) : S₁ = S₂ := eq_of_le_of_finrank_le hle hd.ge -section Subalgebra +end Submodule + +namespace Subalgebra variable {K L : Type*} [Field K] [Ring L] [Algebra K L] {F E : Subalgebra K L} [hfin : FiniteDimensional K E] /-- If a subalgebra is contained in a finite-dimensional subalgebra with the same or smaller dimension, they are equal. -/ -theorem _root_.Subalgebra.eq_of_le_of_finrank_le (h_le : F ≤ E) - (h_finrank : finrank K E ≤ finrank K F) : F = E := +theorem eq_of_le_of_finrank_le (h_le : F ≤ E) (h_finrank : finrank K E ≤ finrank K F) : F = E := haveI : Module.Finite K (Subalgebra.toSubmodule E) := hfin - Subalgebra.toSubmodule_injective <| FiniteDimensional.eq_of_le_of_finrank_le h_le h_finrank + toSubmodule_injective <| Submodule.eq_of_le_of_finrank_le h_le h_finrank /-- If a subalgebra is contained in a finite-dimensional subalgebra with the same dimension, they are equal. -/ -theorem _root_.Subalgebra.eq_of_le_of_finrank_eq (h_le : F ≤ E) - (h_finrank : finrank K F = finrank K E) : F = E := - Subalgebra.eq_of_le_of_finrank_le h_le h_finrank.ge +theorem eq_of_le_of_finrank_eq (h_le : F ≤ E) (h_finrank : finrank K F = finrank K E) : F = E := + eq_of_le_of_finrank_le h_le h_finrank.ge end Subalgebra -end DivisionRing - -end FiniteDimensional - namespace LinearMap -open FiniteDimensional +open Module section DivisionRing @@ -599,7 +592,7 @@ end LinearMap namespace LinearEquiv -open FiniteDimensional +open Module variable [DivisionRing K] [AddCommGroup V] [Module K V] variable [FiniteDimensional K V] @@ -646,14 +639,14 @@ theorem isUnit_iff_range_eq_top [FiniteDimensional K V] (f : V →ₗ[K] V) : end LinearMap -open Module FiniteDimensional +open FiniteDimensional Module section variable [DivisionRing K] [AddCommGroup V] [Module K V] theorem finrank_zero_iff_forall_zero [FiniteDimensional K V] : finrank K V = 0 ↔ ∀ x : V, x = 0 := - FiniteDimensional.finrank_zero_iff.trans (subsingleton_iff_forall_eq 0) + Module.finrank_zero_iff.trans (subsingleton_iff_forall_eq 0) /-- If `ι` is an empty type and `V` is zero-dimensional, there is a unique `ι`-indexed basis. -/ noncomputable def basisOfFinrankZero [FiniteDimensional K V] {ι : Type*} [IsEmpty ι] diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean index d0b35b5f4be62..27e40eb067043 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean @@ -26,7 +26,7 @@ variable (R : Type u) (S : Type u') (M : Type v) (N : Type w) open Module.Free (chooseBasis ChooseBasisIndex) -open FiniteDimensional (finrank) +open Module (finrank) section Ring @@ -46,27 +46,27 @@ instance Module.Finite.linearMap [Module.Finite S N] : Module.Finite S (M →ₗ variable [StrongRankCondition R] [StrongRankCondition S] [Module.Free S N] open Cardinal -theorem FiniteDimensional.rank_linearMap : +theorem Module.rank_linearMap : Module.rank S (M →ₗ[R] N) = lift.{w} (Module.rank R M) * lift.{v} (Module.rank S N) := by rw [(linearMapEquivFun R S M N).rank_eq, rank_fun_eq_lift_mul, ← finrank_eq_card_chooseBasisIndex, ← finrank_eq_rank R, lift_natCast] /-- The finrank of `M →ₗ[R] N` as an `S`-module is `(finrank R M) * (finrank S N)`. -/ -theorem FiniteDimensional.finrank_linearMap : +theorem Module.finrank_linearMap : finrank S (M →ₗ[R] N) = finrank R M * finrank S N := by simp_rw [finrank, rank_linearMap, toNat_mul, toNat_lift] variable [Module R S] [SMulCommClass R S S] -theorem FiniteDimensional.rank_linearMap_self : +theorem Module.rank_linearMap_self : Module.rank S (M →ₗ[R] S) = lift.{u'} (Module.rank R M) := by rw [rank_linearMap, rank_self, lift_one, mul_one] -theorem FiniteDimensional.finrank_linearMap_self : finrank S (M →ₗ[R] S) = finrank R M := by +theorem Module.finrank_linearMap_self : finrank S (M →ₗ[R] S) = finrank R M := by rw [finrank_linearMap, finrank_self, mul_one] @[deprecated (since := "2024-01-12")] -alias FiniteDimensional.finrank_linear_map' := FiniteDimensional.finrank_linearMap_self +alias Module.finrank_linear_map' := Module.finrank_linearMap_self end Ring @@ -84,12 +84,12 @@ theorem cardinal_mk_algHom_le_rank : #(M →ₐ[K] L) ≤ lift.{v} (Module.rank convert (linearIndependent_algHom_toLinearMap K M L).cardinal_lift_le_rank · rw [lift_id] · have := Module.nontrivial K L - rw [lift_id, FiniteDimensional.rank_linearMap_self] + rw [lift_id, Module.rank_linearMap_self] theorem card_algHom_le_finrank : Nat.card (M →ₐ[K] L) ≤ finrank K M := by convert toNat_le_toNat (cardinal_mk_algHom_le_rank K M L) ?_ · rw [toNat_lift, finrank] - · rw [lift_lt_aleph0]; have := Module.nontrivial K L; apply rank_lt_aleph0 + · rw [lift_lt_aleph0]; have := Module.nontrivial K L; apply Module.rank_lt_aleph0 end AlgHom diff --git a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean index 44870709f726a..054d145f3347d 100644 --- a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean +++ b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean @@ -114,9 +114,9 @@ noncomputable def quotientEquivDirectSum : theorem finrank_quotient_eq_sum {ι} [Fintype ι] (b : Basis ι R S) [Nontrivial F] [∀ i, Module.Free F (R ⧸ span ({I.smithCoeffs b hI i} : Set R))] [∀ i, Module.Finite F (R ⧸ span ({I.smithCoeffs b hI i} : Set R))] : - FiniteDimensional.finrank F (S ⧸ I) = - ∑ i, FiniteDimensional.finrank F (R ⧸ span ({I.smithCoeffs b hI i} : Set R)) := by + Module.finrank F (S ⧸ I) = + ∑ i, Module.finrank F (R ⧸ span ({I.smithCoeffs b hI i} : Set R)) := by -- slow, and dot notation doesn't work - rw [LinearEquiv.finrank_eq <| quotientEquivDirectSum F b hI, FiniteDimensional.finrank_directSum] + rw [LinearEquiv.finrank_eq <| quotientEquivDirectSum F b hI, Module.finrank_directSum] end Ideal diff --git a/Mathlib/LinearAlgebra/FreeModule/Norm.lean b/Mathlib/LinearAlgebra/FreeModule/Norm.lean index 39bb414d9ce5c..5332fcc328e26 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Norm.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Norm.lean @@ -71,7 +71,7 @@ instance (b : Basis ι F[X] S) {I : Ideal S} (hI : I ≠ ⊥) (i : ι) : `F`-vector space is the degree of the norm of `f` relative to `F[X]`. -/ theorem finrank_quotient_span_eq_natDegree_norm [Algebra F S] [IsScalarTower F F[X] S] (b : Basis ι F[X] S) {f : S} (hf : f ≠ 0) : - FiniteDimensional.finrank F (S ⧸ span ({f} : Set S)) = (Algebra.norm F[X] f).natDegree := by + Module.finrank F (S ⧸ span ({f} : Set S)) = (Algebra.norm F[X] f).natDegree := by haveI := Fintype.ofFinite ι have h := span_singleton_eq_bot.not.2 hf rw [natDegree_eq_of_degree_eq diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Card.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Card.lean index 8d2b791e900da..d7867798bc310 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Card.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Card.lean @@ -27,7 +27,7 @@ variable {K V : Type*} [DivisionRing K] [AddCommGroup V] [Module K V] variable [Fintype K] [Finite V] local notation "q" => Fintype.card K -local notation "n" => FiniteDimensional.finrank K V +local notation "n" => Module.finrank K V attribute [local instance] Fintype.ofFinite in open Fintype in @@ -86,8 +86,8 @@ theorem card_GL_field : rcases Nat.eq_zero_or_pos n with rfl | hn · simp [Nat.card_eq_fintype_card] · rw [Nat.card_congr (equiv_GL_linearindependent n hn), card_linearIndependent, - FiniteDimensional.finrank_fintype_fun_eq_card, Fintype.card_fin] - simp only [FiniteDimensional.finrank_fintype_fun_eq_card, Fintype.card_fin, le_refl] + Module.finrank_fintype_fun_eq_card, Fintype.card_fin] + simp only [Module.finrank_fintype_fun_eq_card, Fintype.card_fin, le_refl] end field diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index 68891b346f5df..549ff52629c7b 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -326,7 +326,7 @@ namespace Orientation variable [Fintype ι] -open FiniteDimensional +open FiniteDimensional Module /-- If the index type has cardinality equal to the finite dimension, any two orientations are equal or negations. -/ diff --git a/Mathlib/LinearAlgebra/Projectivization/Basic.lean b/Mathlib/LinearAlgebra/Projectivization/Basic.lean index fd5d3d6aa6db7..94d4e7925bb43 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Basic.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Basic.lean @@ -78,7 +78,7 @@ theorem rep_nonzero (v : ℙ K V) : v.rep ≠ 0 := @[simp] theorem mk_rep (v : ℙ K V) : mk K v.rep v.rep_nonzero = v := Quotient.out_eq' _ -open FiniteDimensional +open Module /-- Consider an element of the projectivization as a submodule of `V`. -/ protected def submodule (v : ℙ K V) : Submodule K V := diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 7eba713353773..fe2fe2b66dfe2 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -1180,7 +1180,7 @@ theorem exists_bilinForm_self_ne_zero [htwo : Invertible (2 : R)] {B : BilinForm obtain ⟨x, hx⟩ := QuadraticMap.exists_quadraticForm_ne_zero hB₁ exact ⟨x, fun h => hx (Q.associated_eq_self_apply ℕ x ▸ h)⟩ -open FiniteDimensional +open Module variable {V : Type u} {K : Type v} [Field K] [AddCommGroup V] [Module K V] variable [FiniteDimensional K V] @@ -1194,7 +1194,7 @@ theorem exists_orthogonal_basis [hK : Invertible (2 : K)] {B : LinearMap.BilinFo haveI := finrank_pos_iff.1 (hd.symm ▸ Nat.succ_pos d : 0 < finrank K V) -- either the bilinear form is trivial or we can pick a non-null `x` obtain rfl | hB₁ := eq_or_ne B 0 - · let b := FiniteDimensional.finBasis K V + · let b := Module.finBasis K V rw [hd] at b exact ⟨b, fun i j _ => rfl⟩ obtain ⟨x, hx⟩ := exists_bilinForm_self_ne_zero hB₁ hB₂ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean index 3285feacba1f0..8f1a7a9b3e50d 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean @@ -70,7 +70,7 @@ noncomputable def isometryEquivSumSquaresUnits (w : ι → Units ℂ) : the sum of squares, i.e. `weightedSumSquares` with weight `fun (i : ι) => 1`. -/ theorem equivalent_sum_squares {M : Type*} [AddCommGroup M] [Module ℂ M] [FiniteDimensional ℂ M] (Q : QuadraticForm ℂ M) (hQ : (associated (R := ℂ) Q).SeparatingLeft) : - Equivalent Q (weightedSumSquares ℂ (1 : Fin (FiniteDimensional.finrank ℂ M) → ℂ)) := + Equivalent Q (weightedSumSquares ℂ (1 : Fin (Module.finrank ℂ M) → ℂ)) := let ⟨w, ⟨hw₁⟩⟩ := Q.equivalent_weightedSumSquares_units_of_nondegenerate' hQ ⟨hw₁.trans (isometryEquivSumSquaresUnits w)⟩ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean index ff5c68946f2d8..567af287573fc 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean @@ -141,7 +141,7 @@ variable [Field K] [Invertible (2 : K)] [AddCommGroup V] [Module K V] /-- Given an orthogonal basis, a quadratic form is isometrically equivalent with a weighted sum of squares. -/ noncomputable def isometryEquivWeightedSumSquares (Q : QuadraticForm K V) - (v : Basis (Fin (FiniteDimensional.finrank K V)) K V) + (v : Basis (Fin (Module.finrank K V)) K V) (hv₁ : (associated (R := K) Q).IsOrthoᵢ v) : Q.IsometryEquiv (weightedSumSquares K fun i => Q (v i)) := by let iso := Q.isometryEquivBasisRepr v @@ -154,13 +154,13 @@ variable [FiniteDimensional K V] open LinearMap.BilinForm theorem equivalent_weightedSumSquares (Q : QuadraticForm K V) : - ∃ w : Fin (FiniteDimensional.finrank K V) → K, Equivalent Q (weightedSumSquares K w) := + ∃ w : Fin (Module.finrank K V) → K, Equivalent Q (weightedSumSquares K w) := let ⟨v, hv₁⟩ := exists_orthogonal_basis (associated_isSymm _ Q) ⟨_, ⟨Q.isometryEquivWeightedSumSquares v hv₁⟩⟩ theorem equivalent_weightedSumSquares_units_of_nondegenerate' (Q : QuadraticForm K V) (hQ : (associated (R := K) Q).SeparatingLeft) : - ∃ w : Fin (FiniteDimensional.finrank K V) → Kˣ, Equivalent Q (weightedSumSquares K w) := by + ∃ w : Fin (Module.finrank K V) → Kˣ, Equivalent Q (weightedSumSquares K w) := by obtain ⟨v, hv₁⟩ := exists_orthogonal_basis (associated_isSymm K Q) have hv₂ := hv₁.not_isOrtho_basis_self_of_separatingLeft hQ simp_rw [LinearMap.IsOrtho, associated_eq_self_apply] at hv₂ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean index 7ec3a6f4947d8..b9fe9bdf90ae3 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean @@ -53,7 +53,7 @@ noncomputable def isometryEquivSignWeightedSumSquares (w : ι → ℝ) : sum of squares with the weights being ±1, `SignType` version. -/ theorem equivalent_sign_ne_zero_weighted_sum_squared {M : Type*} [AddCommGroup M] [Module ℝ M] [FiniteDimensional ℝ M] (Q : QuadraticForm ℝ M) (hQ : (associated (R := ℝ) Q).SeparatingLeft) : - ∃ w : Fin (FiniteDimensional.finrank ℝ M) → SignType, + ∃ w : Fin (Module.finrank ℝ M) → SignType, (∀ i, w i ≠ 0) ∧ Equivalent Q (weightedSumSquares ℝ fun i ↦ (w i : ℝ)) := let ⟨w, ⟨hw₁⟩⟩ := Q.equivalent_weightedSumSquares_units_of_nondegenerate' hQ ⟨sign ∘ ((↑) : ℝˣ → ℝ) ∘ w, fun i => sign_ne_zero.2 (w i).ne_zero, @@ -63,7 +63,7 @@ theorem equivalent_sign_ne_zero_weighted_sum_squared {M : Type*} [AddCommGroup M sum of squares with the weights being ±1. -/ theorem equivalent_one_neg_one_weighted_sum_squared {M : Type*} [AddCommGroup M] [Module ℝ M] [FiniteDimensional ℝ M] (Q : QuadraticForm ℝ M) (hQ : (associated (R := ℝ) Q).SeparatingLeft) : - ∃ w : Fin (FiniteDimensional.finrank ℝ M) → ℝ, + ∃ w : Fin (Module.finrank ℝ M) → ℝ, (∀ i, w i = -1 ∨ w i = 1) ∧ Equivalent Q (weightedSumSquares ℝ w) := let ⟨w, hw₀, hw⟩ := Q.equivalent_sign_ne_zero_weighted_sum_squared hQ ⟨(w ·), fun i ↦ by cases hi : w i <;> simp_all, hw⟩ @@ -72,7 +72,7 @@ theorem equivalent_one_neg_one_weighted_sum_squared {M : Type*} [AddCommGroup M] sum of squares with the weights being ±1 or 0, `SignType` version. -/ theorem equivalent_signType_weighted_sum_squared {M : Type*} [AddCommGroup M] [Module ℝ M] [FiniteDimensional ℝ M] (Q : QuadraticForm ℝ M) : - ∃ w : Fin (FiniteDimensional.finrank ℝ M) → SignType, + ∃ w : Fin (Module.finrank ℝ M) → SignType, Equivalent Q (weightedSumSquares ℝ fun i ↦ (w i : ℝ)) := let ⟨w, ⟨hw₁⟩⟩ := Q.equivalent_weightedSumSquares ⟨sign ∘ w, ⟨hw₁.trans (isometryEquivSignWeightedSumSquares w)⟩⟩ @@ -81,7 +81,7 @@ theorem equivalent_signType_weighted_sum_squared {M : Type*} [AddCommGroup M] [M sum of squares with the weights being ±1 or 0. -/ theorem equivalent_one_zero_neg_one_weighted_sum_squared {M : Type*} [AddCommGroup M] [Module ℝ M] [FiniteDimensional ℝ M] (Q : QuadraticForm ℝ M) : - ∃ w : Fin (FiniteDimensional.finrank ℝ M) → ℝ, + ∃ w : Fin (Module.finrank ℝ M) → ℝ, (∀ i, w i = -1 ∨ w i = 0 ∨ w i = 1) ∧ Equivalent Q (weightedSumSquares ℝ w) := let ⟨w, hw⟩ := Q.equivalent_signType_weighted_sum_squared ⟨(w ·), fun i ↦ by cases h : w i <;> simp [h], hw⟩ diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index ea02d0bd4374e..0b3d984857b62 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -145,7 +145,7 @@ theorem IsSemisimple.minpoly_squarefree : Squarefree (minpoly K f) := protected theorem IsSemisimple.aeval (p : K[X]) : (aeval f p).IsSemisimple := let R := K[X] ⧸ Ideal.span {minpoly K f} - have : Finite K R := + have : Module.Finite K R := (AdjoinRoot.powerBasis' <| minpoly.monic <| Algebra.IsIntegral.isIntegral f).finite have : IsReduced R := (Ideal.isRadical_iff_quotient_reduced _).mp <| span_minpoly_eq_annihilator K f ▸ hf.annihilator_isRadical @@ -174,9 +174,9 @@ theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin a.IsSemisimple := by let R := K[X] ⧸ Ideal.span {minpoly K f} let S := AdjoinRoot ((minpoly K g).map <| algebraMap K R) - have : Finite K R := + have : Module.Finite K R := (AdjoinRoot.powerBasis' <| minpoly.monic <| Algebra.IsIntegral.isIntegral f).finite - have : Finite R S := + have : Module.Finite R S := (AdjoinRoot.powerBasis' <| (minpoly.monic <| Algebra.IsIntegral.isIntegral g).map _).finite #adaptation_note /-- @@ -187,7 +187,7 @@ theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin -/ set_option maxSynthPendingDepth 2 in have : IsScalarTower K R S := .of_algebraMap_eq fun _ ↦ rfl - have : Finite K S := .trans R S + have : Module.Finite K S := .trans R S have : IsArtinianRing R := .of_finite K R have : IsReduced R := (Ideal.isRadical_iff_quotient_reduced _).mp <| span_minpoly_eq_annihilator K f ▸ hf.annihilator_isRadical diff --git a/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean b/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean index 2ec0b351ea9e9..83b52e58279fd 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean @@ -36,7 +36,7 @@ mainly used in the definition of linearly disjointness. open scoped TensorProduct -open FiniteDimensional +open Module noncomputable section @@ -194,7 +194,7 @@ theorem rank_sup_le_of_free [Module.Free R A] [Module.Free R B] : exact rank_range_le (A.mulMap B).toLinearMap /-- If `A` and `B` are subalgebras of a commutative `R`-algebra `S`, both of them are -free `R`-algebras, then the `FiniteDimensional.finrank` of `A ⊔ B` is less than or equal to +free `R`-algebras, then the `Module.finrank` of `A ⊔ B` is less than or equal to the product of that of `A` and `B`. -/ theorem finrank_sup_le_of_free [Module.Free R A] [Module.Free R B] : finrank R ↥(A ⊔ B) ≤ finrank R A * finrank R B := by @@ -206,7 +206,7 @@ theorem finrank_sup_le_of_free [Module.Free R A] [Module.Free R B] : wlog hA : ¬ Module.Finite R A generalizing A B · have := this B A (fun h' ↦ h h'.symm) (not_and.1 h (of_not_not hA)) rwa [sup_comm, mul_comm] at this - rw [← Module.rank_lt_alpeh0_iff, not_lt] at hA + rw [← Module.rank_lt_aleph0_iff, not_lt] at hA have := LinearMap.rank_le_of_injective _ <| Submodule.inclusion_injective <| show toSubmodule A ≤ toSubmodule (A ⊔ B) by simp rw [show finrank R A = 0 from Cardinal.toNat_apply_of_aleph0_le hA, diff --git a/Mathlib/LinearAlgebra/Trace.lean b/Mathlib/LinearAlgebra/Trace.lean index d8dfa8fb4f30b..81598681a3b50 100644 --- a/Mathlib/LinearAlgebra/Trace.lean +++ b/Mathlib/LinearAlgebra/Trace.lean @@ -25,7 +25,7 @@ universe u v w namespace LinearMap open scoped Matrix -open FiniteDimensional TensorProduct +open Module TensorProduct section diff --git a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean index 76566b1a9cc1e..2b64a512ceb6e 100644 --- a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean +++ b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean @@ -27,7 +27,7 @@ for a general nontrivial normed space. open Set Function Metric MeasurableSpace intervalIntegral open scoped Pointwise ENNReal NNReal -local notation "dim" => FiniteDimensional.finrank ℝ +local notation "dim" => Module.finrank ℝ noncomputable section namespace MeasureTheory @@ -39,7 +39,7 @@ namespace Measure /-- If `μ` is an additive Haar measure on a normed space `E`, then `μ.toSphere` is the measure on the unit sphere in `E` -such that `μ.toSphere s = FiniteDimensional.finrank ℝ E • μ (Set.Ioo (0 : ℝ) 1 • s)`. -/ +such that `μ.toSphere s = Module.finrank ℝ E • μ (Set.Ioo (0 : ℝ) 1 • s)`. -/ def toSphere (μ : Measure E) : Measure (sphere (0 : E) 1) := dim E • ((μ.comap (Subtype.val ∘ (homeomorphUnitSphereProd E).symm)).restrict (univ ×ˢ Iio ⟨1, mem_Ioi.2 one_pos⟩)).fst @@ -106,7 +106,7 @@ instance (n : ℕ) : SigmaFinite (volumeIoiPow n) := /-- The homeomorphism `homeomorphUnitSphereProd E` sends an additive Haar measure `μ` to the product of `μ.toSphere` and `MeasureTheory.Measure.volumeIoiPow (dim E - 1)`, -where `dim E = FiniteDimensional.finrank ℝ E` is the dimension of `E`. -/ +where `dim E = Module.finrank ℝ E` is the dimension of `E`. -/ theorem measurePreserving_homeomorphUnitSphereProd : MeasurePreserving (homeomorphUnitSphereProd E) (μ.comap (↑)) (μ.toSphere.prod (volumeIoiPow (dim E - 1))) := by @@ -119,7 +119,7 @@ theorem measurePreserving_homeomorphUnitSphereProd : fun s hs ↦ forall_mem_range.2 fun r ↦ ?_ have : Ioo (0 : ℝ) r = r.1 • Ioo (0 : ℝ) 1 := by rw [LinearOrderedField.smul_Ioo r.2.out, smul_zero, smul_eq_mul, mul_one] - have hpos : 0 < dim E := FiniteDimensional.finrank_pos + have hpos : 0 < dim E := Module.finrank_pos rw [(Homeomorph.measurableEmbedding _).map_apply, toSphere_apply' _ hs, volumeIoiPow_apply_Iio, comap_subtype_coe_apply (measurableSet_singleton _).compl, toSphere_apply_aux, this, smul_assoc, μ.addHaar_smul_of_nonneg r.2.out.le, Nat.sub_add_cancel hpos, Nat.cast_pred hpos, diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index 60ca4af5a4851..2c3bc40927bd3 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -43,7 +43,7 @@ In particular, this number is bounded by `5 ^ dim` by a straightforward measure universe u -open Metric Set FiniteDimensional MeasureTheory Filter Fin +open Metric Set Module MeasureTheory Filter Fin open scoped ENNReal Topology diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 3b9bc2e1d765d..ce6116b747cb5 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -87,7 +87,7 @@ Change of variables in integrals [Fremlin, *Measure Theory* (volume 2)][fremlin_vol2] -/ -open MeasureTheory MeasureTheory.Measure Metric Filter Set FiniteDimensional Asymptotics +open MeasureTheory MeasureTheory.Measure Metric Filter Set Module Asymptotics TopologicalSpace open scoped NNReal ENNReal Topology Pointwise @@ -1185,7 +1185,7 @@ theorem det_one_smulRight {𝕜 : Type*} [NormedField 𝕜] (v : 𝕜) : Algebra.id.smul_eq_mul, one_mul, ContinuousLinearMap.coe_smul', Pi.smul_apply, mul_one] rw [this, ContinuousLinearMap.det, ContinuousLinearMap.coe_smul, ContinuousLinearMap.one_def, ContinuousLinearMap.coe_id, LinearMap.det_smul, - FiniteDimensional.finrank_self, LinearMap.det_id, pow_one, mul_one] + Module.finrank_self, LinearMap.det_id, pow_one, mul_one] /-- Integrability in the change of variable formula for differentiable functions (one-variable version): if a function `f` is injective and differentiable on a measurable set `s ⊆ ℝ`, then a diff --git a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean index df01cd8cae6c3..4f7644744e85c 100644 --- a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean +++ b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean @@ -36,7 +36,7 @@ Hermann Minkowski. namespace MeasureTheory -open ENNReal FiniteDimensional MeasureTheory MeasureTheory.Measure Set Filter +open ENNReal Module MeasureTheory MeasureTheory.Measure Set Filter open scoped Pointwise NNReal diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index df652367ebf11..845644d23d8fb 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -392,7 +392,7 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_continuou ### Peak functions of the form `x ↦ c ^ dim * φ (c x)` -/ -open FiniteDimensional Bornology +open Module Bornology variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [FiniteDimensional ℝ F] [MeasurableSpace F] [BorelSpace F] {μ : Measure F} [IsAddHaarMeasure μ] diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index cc3daa531971b..71eec8c794d88 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -16,7 +16,7 @@ measure `1` to the parallelepiped spanned by any orthonormal basis, and that it the canonical `volume` from the `MeasureSpace` instance. -/ -open FiniteDimensional MeasureTheory MeasureTheory.Measure Set +open Module MeasureTheory MeasureTheory.Measure Set variable {ι E F : Type*} diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index 29af339ad2552..a4030ed00f9d7 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -18,7 +18,7 @@ open scoped NNReal ENNReal Pointwise Topology open Inv Set Function MeasureTheory.Measure Filter -open FiniteDimensional +open Module namespace MeasureTheory @@ -122,11 +122,11 @@ alias set_integral_comp_smul_of_pos := setIntegral_comp_smul_of_pos theorem integral_comp_mul_left (g : ℝ → F) (a : ℝ) : (∫ x : ℝ, g (a * x)) = |a⁻¹| • ∫ y : ℝ, g y := by - simp_rw [← smul_eq_mul, Measure.integral_comp_smul, FiniteDimensional.finrank_self, pow_one] + simp_rw [← smul_eq_mul, Measure.integral_comp_smul, Module.finrank_self, pow_one] theorem integral_comp_inv_mul_left (g : ℝ → F) (a : ℝ) : (∫ x : ℝ, g (a⁻¹ * x)) = |a| • ∫ y : ℝ, g y := by - simp_rw [← smul_eq_mul, Measure.integral_comp_inv_smul, FiniteDimensional.finrank_self, pow_one] + simp_rw [← smul_eq_mul, Measure.integral_comp_inv_smul, Module.finrank_self, pow_one] theorem integral_comp_mul_right (g : ℝ → F) (a : ℝ) : (∫ x : ℝ, g (x * a)) = |a⁻¹| • ∫ y : ℝ, g y := by diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 18ebba51fd6f0..8bbb057670b05 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -26,7 +26,7 @@ of the basis). -/ -open Set TopologicalSpace MeasureTheory MeasureTheory.Measure FiniteDimensional +open Set TopologicalSpace MeasureTheory MeasureTheory.Measure Module open scoped Pointwise diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index e03f24092e1a1..497ba69408ba9 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -109,7 +109,7 @@ Hausdorff measure, measure, metric measure open scoped NNReal ENNReal Topology -open EMetric Set Function Filter Encodable FiniteDimensional TopologicalSpace +open EMetric Set Function Filter Encodable Module TopologicalSpace noncomputable section diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 54c09f7095a7a..9ede8fdf04384 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -100,7 +100,7 @@ theorem Basis.map_addHaar {ι E F : Type*} [Fintype ι] [NormedAddCommGroup E] [ namespace MeasureTheory -open Measure TopologicalSpace.PositiveCompacts FiniteDimensional +open Measure TopologicalSpace.PositiveCompacts Module /-! ### The Lebesgue measure is a Haar measure on `ℝ` and on `ℝ^ι`. diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index 110086e7b6708..73abc051da088 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -40,7 +40,7 @@ Using these formulas, we compute the volume of the unit balls in several cases. section general_case -open MeasureTheory MeasureTheory.Measure FiniteDimensional ENNReal +open MeasureTheory MeasureTheory.Measure Module ENNReal theorem MeasureTheory.measure_unitBall_eq_integral_div_gamma {E : Type*} {p : ℝ} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] [MeasurableSpace E] @@ -160,7 +160,7 @@ end general_case section LpSpace -open Real Fintype ENNReal FiniteDimensional MeasureTheory MeasureTheory.Measure +open Real Fintype ENNReal Module MeasureTheory MeasureTheory.Measure variable (ι : Type*) [Fintype ι] {p : ℝ} @@ -350,7 +350,7 @@ end EuclideanSpace section InnerProductSpace -open MeasureTheory MeasureTheory.Measure ENNReal Real FiniteDimensional +open MeasureTheory MeasureTheory.Measure ENNReal Real Module variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] [Nontrivial E] diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 20d24352136cb..3c95d15cc13ee 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -58,7 +58,7 @@ included in the `Cyclotomic` locale. -/ -open Polynomial Algebra FiniteDimensional Set +open Polynomial Algebra Module Set universe u v w z diff --git a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean index b3d2a56cf11a2..642d6933fad59 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean @@ -21,7 +21,7 @@ universe u namespace IsCyclotomicExtension.Rat -open NumberField InfinitePlace FiniteDimensional Complex Nat Polynomial +open NumberField InfinitePlace Module Complex Nat Polynomial variable {n : ℕ+} (K : Type u) [Field K] [CharZero K] diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 33bb6cad832cd..a9c07d9408500 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -61,7 +61,7 @@ and only at the "final step", when we need to provide an "explicit" primitive ro -/ -open Polynomial Algebra Finset FiniteDimensional IsCyclotomicExtension Nat PNat Set +open Polynomial Algebra Finset Module IsCyclotomicExtension Nat PNat Set open scoped IntermediateField universe u v w z @@ -182,7 +182,7 @@ least `(lcm p q).totient`. -/ theorem _root_.IsPrimitiveRoot.lcm_totient_le_finrank [FiniteDimensional K L] {p q : ℕ} {x y : L} (hx : IsPrimitiveRoot x p) (hy : IsPrimitiveRoot y q) (hirr : Irreducible (cyclotomic (Nat.lcm p q) K)) : - (Nat.lcm p q).totient ≤ FiniteDimensional.finrank K L := by + (Nat.lcm p q).totient ≤ Module.finrank K L := by rcases Nat.eq_zero_or_pos p with (rfl | hppos) · simp rcases Nat.eq_zero_or_pos q with (rfl | hqpos) @@ -439,7 +439,7 @@ theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot congr · rw [PNat.pow_coe, Nat.pow_minFac, hpri.1.minFac_eq] exact Nat.succ_ne_zero _ - have := FiniteDimensional.finrank_mul_finrank K K⟮η⟯ L + have := Module.finrank_mul_finrank K K⟮η⟯ L rw [IsCyclotomicExtension.finrank L hirr, IsCyclotomicExtension.finrank K⟮η⟯ hirr₁, PNat.pow_coe, PNat.pow_coe, Nat.totient_prime_pow hpri.out (k - s).succ_pos, Nat.totient_prime_pow hpri.out k.succ_pos, mul_comm _ ((p : ℕ) - 1), mul_assoc, diff --git a/Mathlib/NumberTheory/FunctionField.lean b/Mathlib/NumberTheory/FunctionField.lean index 2093382c27cb9..425189bfa3c7b 100644 --- a/Mathlib/NumberTheory/FunctionField.lean +++ b/Mathlib/NumberTheory/FunctionField.lean @@ -69,9 +69,9 @@ theorem functionField_iff (Fqt : Type*) [Field Fqt] [Algebra Fq[X] Fqt] refine IsLocalization.ext (nonZeroDivisors Fq[X]) _ _ ?_ ?_ ?_ ?_ ?_ <;> intros <;> simp only [map_one, map_mul, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] constructor <;> intro h - · let b := FiniteDimensional.finBasis (RatFunc Fq) F + · let b := Module.finBasis (RatFunc Fq) F exact FiniteDimensional.of_fintype_basis (b.mapCoeffs e this) - · let b := FiniteDimensional.finBasis Fqt F + · let b := Module.finBasis Fqt F refine FiniteDimensional.of_fintype_basis (b.mapCoeffs e.symm ?_) intro c x; convert (this (e.symm c) x).symm; simp only [e.apply_symm_apply] diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 7aef3e02c40b2..cd142c881c370 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -307,7 +307,7 @@ theorem mem_span_integralBasis {x : K} : rw [integralBasis, Basis.localizationLocalization_span, LinearMap.mem_range, IsScalarTower.coe_toAlgHom', RingHom.mem_range] -theorem RingOfIntegers.rank : FiniteDimensional.finrank ℤ (𝓞 K) = FiniteDimensional.finrank ℚ K := +theorem RingOfIntegers.rank : Module.finrank ℤ (𝓞 K) = Module.finrank ℚ K := IsIntegralClosure.rank ℤ ℚ K (𝓞 K) end NumberField diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 5970d297810cb..c62e465bc30d5 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -100,7 +100,7 @@ theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ -open Module Fintype FiniteDimensional +open Module Fintype Module /-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ noncomputable def latticeBasis [NumberField K] : @@ -176,7 +176,7 @@ end NumberField.canonicalEmbedding namespace NumberField.mixedEmbedding -open NumberField.InfinitePlace FiniteDimensional Finset +open NumberField.InfinitePlace Module Finset /-- The mixed space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -/ abbrev mixedSpace := @@ -491,7 +491,7 @@ def indexEquiv : (index K) ≃ (K →+* ℂ) := by · exact ⟨Sum.inr ⟨InfinitePlace.mkComplex ⟨φ, hφ⟩, 1⟩, by simp [(embedding_mk_eq φ).resolve_left hw]⟩ · rw [Embeddings.card, ← mixedEmbedding.finrank K, - ← FiniteDimensional.finrank_eq_card_basis (stdBasis K)] + ← Module.finrank_eq_card_basis (stdBasis K)] variable {K} diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index eb0293caf7695..36a06aabe738b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -40,7 +40,7 @@ variable (K : Type*) [Field K] namespace NumberField.mixedEmbedding -open NumberField NumberField.InfinitePlace FiniteDimensional +open NumberField NumberField.InfinitePlace Module section convexBodyLT @@ -450,7 +450,7 @@ end convexBodySum section minkowski open scoped Classical -open MeasureTheory MeasureTheory.Measure FiniteDimensional ZSpan Real Submodule +open MeasureTheory MeasureTheory.Measure Module ZSpan Real Submodule open scoped ENNReal NNReal nonZeroDivisors IntermediateField diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean index 223002432594b..419b0f6e623c0 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean @@ -87,7 +87,7 @@ end UnitSMul noncomputable section logMap -open NumberField.Units NumberField.Units.dirichletUnitTheorem FiniteDimensional +open NumberField.Units NumberField.Units.dirichletUnitTheorem Module variable [NumberField K] {K} diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index 992a53910d440..06ed53e2b5316 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -41,7 +41,7 @@ variable {K} theorem classNumber_eq_one_iff : classNumber K = 1 ↔ IsPrincipalIdealRing (𝓞 K) := card_classGroup_eq_one_iff -open FiniteDimensional NumberField.InfinitePlace +open Module NumberField.InfinitePlace open scoped nonZeroDivisors Real diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 85e0a1770c687..b56a1d4722194 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -32,7 +32,7 @@ number field, discriminant namespace NumberField -open FiniteDimensional NumberField NumberField.InfinitePlace Matrix +open Module NumberField NumberField.InfinitePlace Matrix open scoped Classical Real nonZeroDivisors diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 40c427c9d06f8..c30dda51f0a43 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -38,7 +38,7 @@ namespace NumberField.Embeddings section Fintype -open FiniteDimensional +open Module variable (K : Type*) [Field K] [NumberField K] variable (A : Type*) [Field A] [CharZero A] @@ -55,7 +55,7 @@ theorem card : Fintype.card (K →+* A) = finrank ℚ K := by instance : Nonempty (K →+* A) := by rw [← Fintype.card_pos_iff, NumberField.Embeddings.card K A] - exact FiniteDimensional.finrank_pos + exact Module.finrank_pos end Fintype @@ -78,7 +78,7 @@ end Roots section Bounded -open FiniteDimensional Polynomial Set +open Module Polynomial Set variable {K : Type*} [Field K] [NumberField K] variable {A : Type*} [NormedField A] [IsAlgClosed A] [NormedAlgebra ℚ A] @@ -450,7 +450,7 @@ noncomputable instance NumberField.InfinitePlace.fintype [NumberField K] : Fintype (InfinitePlace K) := Set.fintypeRange _ theorem sum_mult_eq [NumberField K] : - ∑ w : InfinitePlace K, mult w = FiniteDimensional.finrank ℚ K := by + ∑ w : InfinitePlace K, mult w = Module.finrank ℚ K := by rw [← Embeddings.card K ℂ, Fintype.card, Finset.card_eq_sum_ones, ← Finset.univ.sum_fiberwise (fun φ => InfinitePlace.mk φ)] exact Finset.sum_congr rfl @@ -546,7 +546,7 @@ theorem _root_.NumberField.adjoin_eq_top_of_infinitePlace_lt {x : 𝓞 K} {w : I end NumberField -open Fintype FiniteDimensional +open Fintype Module variable (K) @@ -1024,12 +1024,12 @@ lemma IsUnramifiedAtInfinitePlaces_of_odd_card_aut [IsGalois k K] [FiniteDimensi ⟨fun _ ↦ not_not.mp (Nat.not_even_iff_odd.2 h ∘ InfinitePlace.even_card_aut_of_not_isUnramified)⟩ lemma IsUnramifiedAtInfinitePlaces_of_odd_finrank [IsGalois k K] - (h : Odd (FiniteDimensional.finrank k K)) : IsUnramifiedAtInfinitePlaces k K := + (h : Odd (Module.finrank k K)) : IsUnramifiedAtInfinitePlaces k K := ⟨fun _ ↦ not_not.mp (Nat.not_even_iff_odd.2 h ∘ InfinitePlace.even_finrank_of_not_isUnramified)⟩ variable (k K) -open FiniteDimensional in +open Module in lemma IsUnramifiedAtInfinitePlaces.card_infinitePlace [NumberField k] [NumberField K] [IsGalois k K] [IsUnramifiedAtInfinitePlaces k K] : Fintype.card (InfinitePlace K) = Fintype.card (InfinitePlace k) * finrank k K := by diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean index 67fc4926a8307..0226ce9f8e5d5 100644 --- a/Mathlib/NumberTheory/NumberField/EquivReindex.lean +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -21,7 +21,7 @@ namespace NumberField noncomputable section -open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset +open Module.Free Module canonicalEmbedding Matrix Finset /-- 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`. -/ diff --git a/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean b/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean index 47f93f496cc3b..1444de4b7e903 100644 --- a/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean +++ b/Mathlib/NumberTheory/NumberField/FractionalIdeal.lean @@ -89,7 +89,7 @@ theorem mem_span_basisOfFractionalIdeal {I : (FractionalIdeal (𝓞 K)⁰ K)ˣ} rw [basisOfFractionalIdeal, (fractionalIdealBasis K I.1).ofIsLocalizedModule_span ℚ ℤ⁰ _] simp -open FiniteDimensional in +open Module in theorem fractionalIdeal_rank (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : finrank ℤ I = finrank ℤ (𝓞 K) := by rw [finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index 5d0cc0062fc60..d56eb4d18c668 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -27,7 +27,7 @@ namespace NumberField noncomputable section -open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset +open Module.Free Module canonicalEmbedding Matrix Finset attribute [local instance] Matrix.seminormedAddCommGroup @@ -62,7 +62,7 @@ noncomputable section variable (K) -open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset +open Module.Free Module canonicalEmbedding Matrix Finset attribute [local instance] Matrix.seminormedAddCommGroup diff --git a/Mathlib/NumberTheory/NumberField/Norm.lean b/Mathlib/NumberTheory/NumberField/Norm.lean index 2b529a1d239bb..e8964a687a66b 100644 --- a/Mathlib/NumberTheory/NumberField/Norm.lean +++ b/Mathlib/NumberTheory/NumberField/Norm.lean @@ -22,7 +22,7 @@ rings of integers. open scoped NumberField -open Finset NumberField Algebra FiniteDimensional +open Finset NumberField Algebra Module section Rat diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 636189cd1f51f..699d1221574f7 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -354,7 +354,7 @@ section statements variable [NumberField K] open scoped Classical -open dirichletUnitTheorem FiniteDimensional +open dirichletUnitTheorem Module /-- The unit rank of the number field `K`, it is equal to `card (InfinitePlace K) - 1`. -/ def rank : ℕ := Fintype.card (InfinitePlace K) - 1 @@ -462,13 +462,13 @@ instance : Monoid.FG (𝓞 K)ˣ := by infer_instance theorem rank_modTorsion : - FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by + Module.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by rw [← LinearEquiv.finrank_eq (logEmbeddingEquiv K).symm, unitLattice_rank] /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq <| by - rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, rank_modTorsion, Fintype.card_fin]) + rw [← Module.finrank_eq_card_chooseBasisIndex, rank_modTorsion, Fintype.card_fin]) /-- The basis of the `unitLattice` obtained by mapping `basisModTorsion` via `logEmbedding`. -/ def basisUnitLattice : Basis (Fin (rank K)) ℤ (unitLattice K) := diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index c059ce200afdd..c84ce456063d4 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -47,7 +47,7 @@ variable {R : Type u} [CommRing R] variable {S : Type v} [CommRing S] (f : R →+* S) variable (p : Ideal R) (P : Ideal S) -open FiniteDimensional +open Module open UniqueFactorizationMonoid @@ -699,7 +699,7 @@ instance Factors.finiteDimensional_quotient [IsNoetherian R S] [p.IsMaximal] theorem Factors.inertiaDeg_ne_zero [IsNoetherian R S] [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : inertiaDeg (algebraMap R S) p P ≠ 0 := by - rw [inertiaDeg_algebraMap]; exact (FiniteDimensional.finrank_pos_iff.mpr inferInstance).ne' + rw [inertiaDeg_algebraMap]; exact (Module.finrank_pos_iff.mpr inferInstance).ne' instance Factors.finiteDimensional_quotient_pow [IsNoetherian R S] [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index e40c61bfaf1ea..b629a2dc92e8b 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -32,7 +32,7 @@ noncomputable section universe u -open CategoryTheory LinearMap CategoryTheory.MonoidalCategory Representation FiniteDimensional +open CategoryTheory LinearMap CategoryTheory.MonoidalCategory Representation Module variable {k : Type u} [Field k] @@ -51,7 +51,7 @@ theorem char_mul_comm (V : FDRep k G) (g : G) (h : G) : V.character (h * g) = V.character (g * h) := by simp only [trace_mul_comm, character, map_mul] @[simp] -theorem char_one (V : FDRep k G) : V.character 1 = FiniteDimensional.finrank k V := by +theorem char_one (V : FDRep k G) : V.character 1 = Module.finrank k V := by simp only [character, map_one, trace_one] /-- The character is multiplicative under the tensor product. -/ diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index b19b49f0e400b..1f3f8535f0114 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -114,7 +114,7 @@ example : MonoidalPreadditive (FDRep k G) := by infer_instance example : MonoidalLinear k (FDRep k G) := by infer_instance -open FiniteDimensional +open Module open scoped Classical diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index f8aba734476d2..86796148fbcf3 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -198,13 +198,13 @@ theorem IsIntegralClosure.module_free [NoZeroSMulDivisors A L] [IsPrincipalIdeal and `L` has no zero smul divisors by `A`, the `A`-rank of the integral closure `C` of `A` in `L` is equal to the `K`-rank of `L`. -/ theorem IsIntegralClosure.rank [IsPrincipalIdealRing A] [NoZeroSMulDivisors A L] : - FiniteDimensional.finrank A C = FiniteDimensional.finrank K L := by + Module.finrank A C = Module.finrank K L := by haveI : Module.Free A C := IsIntegralClosure.module_free A K L C haveI : IsNoetherian A C := IsIntegralClosure.isNoetherian A K L C haveI : IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := IsIntegralClosure.isLocalization A K L C let b := Basis.localizationLocalization K A⁰ L (Module.Free.chooseBasis A C) - rw [FiniteDimensional.finrank_eq_card_chooseBasisIndex, FiniteDimensional.finrank_eq_card_basis b] + rw [Module.finrank_eq_card_chooseBasisIndex, Module.finrank_eq_card_basis b] variable {A K} diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index b3171e5b848cd..5957f090f0b89 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -29,7 +29,7 @@ variable (R : Type*) [CommRing R] (K : Type*) [Field K] [Algebra R K] [IsFractio open scoped Multiplicative -open LocalRing FiniteDimensional +open LocalRing Module theorem exists_maximalIdeal_pow_eq_of_principal [IsNoetherianRing R] [LocalRing R] [IsDomain R] (h' : (maximalIdeal R).IsPrincipal) (I : Ideal R) (hI : I ≠ ⊥) : diff --git a/Mathlib/RingTheory/Discriminant.lean b/Mathlib/RingTheory/Discriminant.lean index f6285f4b7aa27..f3faffa03b668 100644 --- a/Mathlib/RingTheory/Discriminant.lean +++ b/Mathlib/RingTheory/Discriminant.lean @@ -50,7 +50,7 @@ universe u v w z open scoped Matrix -open Matrix FiniteDimensional Fintype Polynomial Finset IntermediateField +open Matrix Module Fintype Polynomial Finset IntermediateField namespace Algebra @@ -190,7 +190,7 @@ theorem discr_powerBasis_eq_prod'' [Algebra.IsSeparable K L] (e : Fin pb.dim ≃ have h₂ : 2 ∣ pb.dim * (pb.dim - 1) := pb.dim.even_mul_pred_self.two_dvd have hne : ((2 : ℕ) : ℚ) ≠ 0 := by simp have hle : 1 ≤ pb.dim := by - rw [← hn, Nat.one_le_iff_ne_zero, ← zero_lt_iff, FiniteDimensional.finrank_pos_iff] + rw [← hn, Nat.one_le_iff_ne_zero, ← zero_lt_iff, Module.finrank_pos_iff] infer_instance rw [hn, Nat.cast_div h₂ hne, Nat.cast_mul, Nat.cast_sub hle] field_simp diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index b0366a10493a5..349a1ceec0456 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -45,7 +45,7 @@ section Algebra -- see Note [lower instance priority] instance (priority := 100) finiteType {R : Type*} (A : Type*) [CommSemiring R] [Semiring A] - [Algebra R A] [hRA : Finite R A] : Algebra.FiniteType R A := + [Algebra R A] [hRA : Module.Finite R A] : Algebra.FiniteType R A := ⟨Subalgebra.fg_of_submodule_fg hRA.1⟩ end Algebra @@ -746,7 +746,7 @@ This is similar to `IsNoetherian.injective_of_surjective_endomorphism` but only commutative case, but does not use a Noetherian hypothesis. -/ @[deprecated OrzechProperty.injective_of_surjective_endomorphism (since := "2024-05-30")] theorem Module.Finite.injective_of_surjective_endomorphism {R : Type*} [CommRing R] {M : Type*} - [AddCommGroup M] [Module R M] [Finite R M] (f : M →ₗ[R] M) + [AddCommGroup M] [Module R M] [Module.Finite R M] (f : M →ₗ[R] M) (f_surj : Function.Surjective f) : Function.Injective f := OrzechProperty.injective_of_surjective_endomorphism f f_surj diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index b9e7cb2a2c070..5d61d585d9c03 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -501,7 +501,7 @@ section ModuleAndAlgebra variable (R A B M N : Type*) /-- A module over a semiring is `Finite` if it is finitely generated as a module. -/ -class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where +protected class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where out : (⊤ : Submodule R M).FG attribute [inherit_doc Module.Finite] Module.Finite.out @@ -511,7 +511,7 @@ namespace Module variable [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] theorem finite_def {R M} [Semiring R] [AddCommMonoid M] [Module R M] : - Finite R M ↔ (⊤ : Submodule R M).FG := + Module.Finite R M ↔ (⊤ : Submodule R M).FG := ⟨fun h => h.1, fun h => ⟨h⟩⟩ namespace Finite @@ -529,46 +529,47 @@ theorem iff_addGroup_fg {G : Type*} [AddCommGroup G] : Module.Finite ℤ G ↔ A variable {R M N} /-- See also `Module.Finite.exists_fin'`. -/ -theorem exists_fin [Finite R M] : ∃ (n : ℕ) (s : Fin n → M), Submodule.span R (range s) = ⊤ := +lemma exists_fin [Module.Finite R M] : ∃ (n : ℕ) (s : Fin n → M), Submodule.span R (range s) = ⊤ := Submodule.fg_iff_exists_fin_generating_family.mp out variable (R M) in -lemma exists_fin' [Finite R M] : ∃ (n : ℕ) (f : (Fin n → R) →ₗ[R] M), Surjective f := by +lemma exists_fin' [Module.Finite R M] : ∃ (n : ℕ) (f : (Fin n → R) →ₗ[R] M), Surjective f := by have ⟨n, s, hs⟩ := exists_fin (R := R) (M := M) refine ⟨n, Basis.constr (Pi.basisFun R _) ℕ s, ?_⟩ rw [← LinearMap.range_eq_top, Basis.constr_range, hs] -theorem of_surjective [hM : Finite R M] (f : M →ₗ[R] N) (hf : Surjective f) : Finite R N := +theorem of_surjective [hM : Module.Finite R M] (f : M →ₗ[R] N) (hf : Surjective f) : + Module.Finite R N := ⟨by rw [← LinearMap.range_eq_top.2 hf, ← Submodule.map_top] exact hM.1.map f⟩ instance quotient (R) {A M} [Semiring R] [AddCommGroup M] [Ring A] [Module A M] [Module R M] - [SMul R A] [IsScalarTower R A M] [Finite R M] - (N : Submodule A M) : Finite R (M ⧸ N) := + [SMul R A] [IsScalarTower R A M] [Module.Finite R M] + (N : Submodule A M) : Module.Finite R (M ⧸ N) := Module.Finite.of_surjective (N.mkQ.restrictScalars R) N.mkQ_surjective /-- The range of a linear map from a finite module is finite. -/ -instance range {F : Type*} [FunLike F M N] [SemilinearMapClass F (RingHom.id R) M N] [Finite R M] - (f : F) : Finite R (LinearMap.range f) := +instance range {F : Type*} [FunLike F M N] [SemilinearMapClass F (RingHom.id R) M N] + [Module.Finite R M] (f : F) : Module.Finite R (LinearMap.range f) := of_surjective (SemilinearMapClass.semilinearMap f).rangeRestrict fun ⟨_, y, hy⟩ => ⟨y, Subtype.ext hy⟩ /-- Pushforwards of finite submodules are finite. -/ -instance map (p : Submodule R M) [Finite R p] (f : M →ₗ[R] N) : Finite R (p.map f) := +instance map (p : Submodule R M) [Module.Finite R p] (f : M →ₗ[R] N) : Module.Finite R (p.map f) := of_surjective (f.restrict fun _ => Submodule.mem_map_of_mem) fun ⟨_, _, hy, hy'⟩ => ⟨⟨_, hy⟩, Subtype.ext hy'⟩ variable (R) -instance self : Finite R R := +instance self : Module.Finite R R := ⟨⟨{1}, by simpa only [Finset.coe_singleton] using Ideal.span_singleton_one⟩⟩ variable (M) theorem of_restrictScalars_finite (R A M : Type*) [CommSemiring R] [Semiring A] [AddCommMonoid M] - [Module R M] [Module A M] [Algebra R A] [IsScalarTower R A M] [hM : Finite R M] : - Finite A M := by + [Module R M] [Module A M] [Algebra R A] [IsScalarTower R A M] [hM : Module.Finite R M] : + Module.Finite A M := by rw [finite_def, Submodule.fg_def] at hM ⊢ obtain ⟨S, hSfin, hSgen⟩ := hM refine ⟨S, hSfin, eq_top_iff.2 ?_⟩ @@ -578,24 +579,24 @@ theorem of_restrictScalars_finite (R A M : Type*) [CommSemiring R] [Semiring A] variable {R M} -instance prod [hM : Finite R M] [hN : Finite R N] : Finite R (M × N) := +instance prod [hM : Module.Finite R M] [hN : Module.Finite R N] : Module.Finite R (M × N) := ⟨by rw [← Submodule.prod_top] exact hM.1.prod hN.1⟩ instance pi {ι : Type*} {M : ι → Type*} [_root_.Finite ι] [∀ i, AddCommMonoid (M i)] - [∀ i, Module R (M i)] [h : ∀ i, Finite R (M i)] : Finite R (∀ i, M i) := + [∀ i, Module R (M i)] [h : ∀ i, Module.Finite R (M i)] : Module.Finite R (∀ i, M i) := ⟨by rw [← Submodule.pi_top] exact Submodule.fg_pi fun i => (h i).1⟩ -theorem equiv [Finite R M] (e : M ≃ₗ[R] N) : Finite R N := +theorem equiv [Module.Finite R M] (e : M ≃ₗ[R] N) : Module.Finite R N := of_surjective (e : M →ₗ[R] N) e.surjective -theorem equiv_iff (e : M ≃ₗ[R] N) : Finite R M ↔ Finite R N := +theorem equiv_iff (e : M ≃ₗ[R] N) : Module.Finite R M ↔ Module.Finite R N := ⟨fun _ ↦ equiv e, fun _ ↦ equiv e.symm⟩ -instance ulift [Finite R M] : Finite R (ULift M) := equiv ULift.moduleEquiv.symm +instance ulift [Module.Finite R M] : Module.Finite R (ULift M) := equiv ULift.moduleEquiv.symm theorem iff_fg {N : Submodule R M} : Module.Finite R N ↔ N.FG := Module.finite_def.trans (fg_top _) @@ -603,7 +604,7 @@ variable (R M) instance bot : Module.Finite R (⊥ : Submodule R M) := iff_fg.mpr fg_bot -instance top [Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr out +instance top [Module.Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr out variable {M} @@ -642,7 +643,7 @@ section Algebra theorem trans {R : Type*} (A M : Type*) [Semiring R] [Semiring A] [Module R A] [AddCommMonoid M] [Module R M] [Module A M] [IsScalarTower R A M] : - ∀ [Finite R A] [Finite A M], Finite R M + ∀ [Module.Finite R A] [Module.Finite A M], Module.Finite R M | ⟨⟨s, hs⟩⟩, ⟨⟨t, ht⟩⟩ => ⟨Submodule.fg_def.2 ⟨Set.image2 (· • ·) (↑s : Set A) (↑t : Set M), diff --git a/Mathlib/RingTheory/Flat/EquationalCriterion.lean b/Mathlib/RingTheory/Flat/EquationalCriterion.lean index e90cd0987d45c..763d9dc2c83db 100644 --- a/Mathlib/RingTheory/Flat/EquationalCriterion.lean +++ b/Mathlib/RingTheory/Flat/EquationalCriterion.lean @@ -120,8 +120,8 @@ theorem tfae_equational_criterion : List.TFAE [ ∀ {ι : Type u} [Fintype ι] {f : ι →₀ R} {x : (ι →₀ R) →ₗ[R] M}, x f = 0 → ∃ (κ : Type u) (_ : Fintype κ) (a : (ι →₀ R) →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), x = y ∘ₗ a ∧ a f = 0, - ∀ {N : Type u} [AddCommGroup N] [Module R N] [Free R N] [Finite R N] {f : N} {x : N →ₗ[R] M}, - x f = 0 → + ∀ {N : Type u} [AddCommGroup N] [Module R N] [Free R N] [Module.Finite R N] {f : N} + {x : N →ₗ[R] M}, x f = 0 → ∃ (κ : Type u) (_ : Fintype κ) (a : N →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), x = y ∘ₗ a ∧ a f = 0] := by classical @@ -243,8 +243,8 @@ Let $M$ be a flat module over a commutative ring $R$. Let $N$ be a finite free m let $f \in N$, and let $x \colon N \to M$ be a homomorphism such that $x(f) = 0$. Then there exist a finite index type $\kappa$ and module homomorphisms $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ -theorem exists_factorization_of_apply_eq_zero_of_free [Flat R M] {N : Type u} - [AddCommGroup N] [Module R N] [Free R N] [Finite R N] {f : N} {x : N →ₗ[R] M} (h : x f = 0) : +theorem exists_factorization_of_apply_eq_zero_of_free [Flat R M] {N : Type u} [AddCommGroup N] + [Module R N] [Free R N] [Module.Finite R N] {f : N} {x : N →ₗ[R] M} (h : x f = 0) : ∃ (κ : Type u) (_ : Fintype κ) (a : N →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), x = y ∘ₗ a ∧ a f = 0 := by exact ((tfae_equational_criterion R M).out 0 5 rfl rfl).mp ‹Flat R M› h @@ -254,8 +254,8 @@ free, and let $f \colon K \to N$ and $x \colon N \to M$ be homomorphisms such th $x \circ f = 0$. Then there exist a finite index type $\kappa$ and module homomorphisms $a \colon N \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a \circ f = 0$. -/ -theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} - [AddCommGroup K] [Module R K] [Finite R K] [AddCommGroup N] [Module R N] [Free R N] [Finite R N] +theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} [AddCommGroup K] + [Module R K] [Module.Finite R K] [AddCommGroup N] [Module R N] [Free R N] [Module.Finite R N] {f : K →ₗ[R] N} {x : N →ₗ[R] M} (h : x ∘ₗ f = 0) : ∃ (κ : Type u) (_ : Fintype κ) (a : N →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), x = y ∘ₗ a ∧ a ∘ₗ f = 0 := by @@ -283,7 +283,7 @@ theorem exists_factorization_of_isFinitelyPresented [Flat R M] {P : Type u} [Add ∃ (κ : Type u) (_ : Fintype κ) (h₂ : P →ₗ[R] (κ →₀ R)) (h₃ : (κ →₀ R) →ₗ[R] M), h₁ = h₃ ∘ₗ h₂ := by obtain ⟨L, _, _, K, ϕ, _, _, hK⟩ := FinitePresentation.equiv_quotient R P - haveI : Finite R ↥K := Module.Finite.iff_fg.mpr hK + haveI : Module.Finite R ↥K := Module.Finite.iff_fg.mpr hK have : (h₁ ∘ₗ ϕ.symm ∘ₗ K.mkQ) ∘ₗ K.subtype = 0 := by simp_rw [comp_assoc, (LinearMap.exact_subtype_mkQ K).linearMap_comp_eq_zero, comp_zero] obtain ⟨κ, hκ, a, y, hay, ha⟩ := exists_factorization_of_comp_eq_zero_of_free this diff --git a/Mathlib/RingTheory/Ideal/Cotangent.lean b/Mathlib/RingTheory/Ideal/Cotangent.lean index 948ea0ff4843f..9ce01fea0e26b 100644 --- a/Mathlib/RingTheory/Ideal/Cotangent.lean +++ b/Mathlib/RingTheory/Ideal/Cotangent.lean @@ -251,7 +251,7 @@ lemma CotangentSpace.span_image_eq_top_iff [IsNoetherianRing R] {s : Set (maxima · simp only [Ideal.toCotangent_apply, Submodule.restrictScalars_top, Submodule.map_span] · exact Ideal.Quotient.mk_surjective -open FiniteDimensional +open Module lemma finrank_cotangentSpace_eq_zero_iff [IsNoetherianRing R] : finrank (ResidueField R) (CotangentSpace R) = 0 ↔ IsField R := by diff --git a/Mathlib/RingTheory/LittleWedderburn.lean b/Mathlib/RingTheory/LittleWedderburn.lean index 24f9b1cce5f40..42f44f6d9f529 100644 --- a/Mathlib/RingTheory/LittleWedderburn.lean +++ b/Mathlib/RingTheory/LittleWedderburn.lean @@ -48,7 +48,7 @@ private def InductionHyp : Prop := namespace InductionHyp -open FiniteDimensional Polynomial +open Module Polynomial variable {D} diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 06103c98b607e..838c66b8f3863 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -171,12 +171,12 @@ theorem free_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] refine ⟨?_, this⟩ rw [← LinearMap.ker_eq_bot (M := k ⊗[R] (I →₀ R)) (f := i.baseChange k), ← Submodule.finrank_eq_zero (R := k) (M := k ⊗[R] (I →₀ R)), - ← Nat.add_right_inj (n := FiniteDimensional.finrank k (LinearMap.range <| i.baseChange k)), + ← Nat.add_right_inj (n := Module.finrank k (LinearMap.range <| i.baseChange k)), LinearMap.finrank_range_add_finrank_ker (V := k ⊗[R] (I →₀ R)), LinearMap.range_eq_top.mpr this, finrank_top] - simp only [FiniteDimensional.finrank_tensorProduct, FiniteDimensional.finrank_self, - FiniteDimensional.finrank_finsupp_self, one_mul, add_zero] - rw [FiniteDimensional.finrank_eq_card_chooseBasisIndex] + simp only [Module.finrank_tensorProduct, Module.finrank_self, + Module.finrank_finsupp_self, one_mul, add_zero] + rw [Module.finrank_eq_card_chooseBasisIndex] -- On the other hand, `m ⊗ M → M` injective => `Tor₁(k, M) = 0` => `k ⊗ ker(i) → kᴵ` injective. have := @lTensor_injective_of_exact_of_exact_of_rTensor_injective (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index f3e0b396a8eb4..3cf506b139773 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -145,17 +145,17 @@ variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R variable (R M) -- see Note [lower instance priority] -instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Finite R M := +instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := ⟨IsNoetherian.noetherian ⊤⟩ instance {R₁ S : Type*} [CommSemiring R₁] [Semiring S] [Algebra R₁ S] - [IsNoetherian R₁ S] (I : Ideal S) : Finite R₁ I := + [IsNoetherian R₁ S] (I : Ideal S) : Module.Finite R₁ I := IsNoetherian.finite R₁ ((I : Submodule S S).restrictScalars R₁) variable {R M} theorem Finite.of_injective [IsNoetherian R N] (f : M →ₗ[R] N) (hf : Function.Injective f) : - Finite R M := + Module.Finite R M := ⟨fg_of_injective f hf⟩ end Module diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index 142c31df3e1f0..e58b1fc23e208 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -46,7 +46,7 @@ variable {K L F : Type*} [Field K] [Field L] [Field F] variable [Algebra K L] [Algebra K F] variable {ι : Type w} -open FiniteDimensional +open Module open LinearMap @@ -153,7 +153,7 @@ theorem _root_.IntermediateField.AdjoinSimple.norm_gen_eq_one {x : L} (hx : ¬Is contrapose! hx obtain ⟨s, ⟨b⟩⟩ := hx refine .of_mem_of_fg K⟮x⟯.toSubalgebra ?_ x ?_ - · exact (Submodule.fg_iff_finiteDimensional _).mpr (of_fintype_basis b) + · exact (Submodule.fg_iff_finiteDimensional _).mpr (.of_fintype_basis b) · exact IntermediateField.subset_adjoin K _ (Set.mem_singleton x) theorem _root_.IntermediateField.AdjoinSimple.norm_gen_eq_prod_roots (x : L) diff --git a/Mathlib/RingTheory/Norm/Defs.lean b/Mathlib/RingTheory/Norm/Defs.lean index 6f1c21fae5c00..b7f2ce47f5c97 100644 --- a/Mathlib/RingTheory/Norm/Defs.lean +++ b/Mathlib/RingTheory/Norm/Defs.lean @@ -41,7 +41,7 @@ variable {K L F : Type*} [Field K] [Field L] [Field F] variable [Algebra K L] [Algebra K F] variable {ι : Type w} -open FiniteDimensional +open Module open LinearMap diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index a7f446664bdad..6365d391b724f 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -137,7 +137,7 @@ theorem dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt {B : Pow letI := B.finite let P := minpoly R B.gen obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' - have finrank_K_L : FiniteDimensional.finrank K L = B.dim := B.finrank + have finrank_K_L : Module.finrank K L = B.dim := B.finrank have deg_K_P : (minpoly K B.gen).natDegree = B.dim := B.natDegree_minpoly have deg_R_P : P.natDegree = B.dim := by rw [← deg_K_P, minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index b0b9c2d3d7536..a9cf86b0049c1 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -19,7 +19,7 @@ gives a `PowerBasis` structure generated by `x`. * `PowerBasis R A`: a structure containing an `x` and an `n` such that `1, x, ..., x^n` is a basis for the `R`-algebra `A` (viewed as an `R`-module). -* `finrank (hf : f ≠ 0) : FiniteDimensional.finrank K (AdjoinRoot f) = f.natDegree`, +* `finrank (hf : f ≠ 0) : Module.finrank K (AdjoinRoot f) = f.natDegree`, the dimension of `AdjoinRoot f` equals the degree of `f` * `PowerBasis.lift (pb : PowerBasis R S)`: if `y : S'` satisfies the same @@ -77,8 +77,8 @@ theorem finite (pb : PowerBasis R S) : Module.Finite R S := .of_basis pb.basis @[deprecated (since := "2024-03-05")] alias finiteDimensional := PowerBasis.finite theorem finrank [StrongRankCondition R] (pb : PowerBasis R S) : - FiniteDimensional.finrank R S = pb.dim := by - rw [FiniteDimensional.finrank_eq_card_basis pb.basis, Fintype.card_fin] + Module.finrank R S = pb.dim := by + rw [Module.finrank_eq_card_basis pb.basis, Fintype.card_fin] theorem mem_span_pow' {x y : S} {d : ℕ} : y ∈ Submodule.span R (Set.range fun i : Fin d => x ^ (i : ℕ)) ↔ diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 5bc001188d6d1..86e5d7e807588 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -163,7 +163,7 @@ theorem isSimpleModule_self_iff_isUnit : exact ⟨⟨x, y, left_inv_eq_right_inv hzy hyx ▸ hzy, hyx⟩, rfl⟩ theorem isSimpleModule_iff_finrank_eq_one {R} [DivisionRing R] [Module R M] : - IsSimpleModule R M ↔ FiniteDimensional.finrank R M = 1 := + IsSimpleModule R M ↔ Module.finrank R M = 1 := ⟨fun h ↦ have := h.nontrivial; have ⟨v, hv⟩ := exists_ne (0 : M) (finrank_eq_one_iff_of_nonzero' v hv).mpr (IsSimpleModule.toSpanSingleton_surjective R hv), is_simple_module_of_finrank_eq_one⟩ diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 6690ed277ff91..e06b0141ff923 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -47,7 +47,7 @@ variable [Algebra R S] [Algebra R T] variable {K L : Type*} [Field K] [Field L] [Algebra K L] variable {ι κ : Type w} [Fintype ι] -open FiniteDimensional +open Module open LinearMap (BilinForm) open LinearMap @@ -435,7 +435,7 @@ variable (K L) theorem traceForm_nondegenerate [FiniteDimensional K L] [Algebra.IsSeparable K L] : (traceForm K L).Nondegenerate := BilinForm.nondegenerate_of_det_ne_zero (traceForm K L) _ - (det_traceForm_ne_zero (FiniteDimensional.finBasis K L)) + (det_traceForm_ne_zero (Module.finBasis K L)) theorem Algebra.trace_ne_zero [FiniteDimensional K L] [Algebra.IsSeparable K L] : Algebra.trace K L ≠ 0 := by diff --git a/Mathlib/RingTheory/Trace/Defs.lean b/Mathlib/RingTheory/Trace/Defs.lean index 86218f0b3ecff..8d961bf7c82e6 100644 --- a/Mathlib/RingTheory/Trace/Defs.lean +++ b/Mathlib/RingTheory/Trace/Defs.lean @@ -48,7 +48,7 @@ variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] variable [Algebra R S] [Algebra R T] variable {ι κ : Type w} [Fintype ι] -open FiniteDimensional +open Module open LinearMap (BilinForm) open LinearMap diff --git a/Mathlib/RingTheory/Valuation/Minpoly.lean b/Mathlib/RingTheory/Valuation/Minpoly.lean index 3448962da4e70..00094cc933c86 100644 --- a/Mathlib/RingTheory/Valuation/Minpoly.lean +++ b/Mathlib/RingTheory/Valuation/Minpoly.lean @@ -21,7 +21,7 @@ Let `K` be a field with a valuation `v` and let `L` be a field extension of `K`. is helpful for defining the valuation on `L` inducing `v`. -/ -open FiniteDimensional minpoly Polynomial +open Module minpoly Polynomial variable {K : Type*} [Field K] {Γ₀ : Type*} [LinearOrderedCommGroupWithZero Γ₀] (v : Valuation K Γ₀) (L : Type*) [Field L] [Algebra K L] diff --git a/Mathlib/RingTheory/WittVector/Isocrystal.lean b/Mathlib/RingTheory/WittVector/Isocrystal.lean index 5ba7feb02f7c6..0f5d266cc6c93 100644 --- a/Mathlib/RingTheory/WittVector/Isocrystal.lean +++ b/Mathlib/RingTheory/WittVector/Isocrystal.lean @@ -54,7 +54,7 @@ This file introduces notation in the locale `Isocrystal`. noncomputable section -open FiniteDimensional +open Module namespace WittVector @@ -181,7 +181,7 @@ admits an isomorphism to one of the standard (indexed by `m : ℤ`) one-dimensio theorem isocrystal_classification (k : Type*) [Field k] [IsAlgClosed k] [CharP k p] (V : Type*) [AddCommGroup V] [Isocrystal p k V] (h_dim : finrank K(p, k) V = 1) : ∃ m : ℤ, Nonempty (StandardOneDimIsocrystal p k m ≃ᶠⁱ[p, k] V) := by - haveI : Nontrivial V := FiniteDimensional.nontrivial_of_finrank_eq_succ h_dim + haveI : Nontrivial V := Module.nontrivial_of_finrank_eq_succ h_dim obtain ⟨x, hx⟩ : ∃ x : V, x ≠ 0 := exists_ne 0 have : Φ(p, k) x ≠ 0 := by simpa only [map_zero] using Φ(p, k).injective.ne hx obtain ⟨a, ha, hax⟩ : ∃ a : K(p, k), a ≠ 0 ∧ Φ(p, k) x = a • x := by diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 31b45d485cdb9..3c7e1401aa012 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -47,7 +47,7 @@ universe u v w x noncomputable section -open Set FiniteDimensional TopologicalSpace Filter +open Filter Module Set TopologicalSpace section Field @@ -197,7 +197,7 @@ private theorem continuous_equivFun_basis_aux [T2Space E] {ι : Type v} [Fintype induction' hn : Fintype.card ι with n IH generalizing ι E · rw [Fintype.card_eq_zero_iff] at hn exact continuous_of_const fun x y => funext hn.elim - · haveI : FiniteDimensional 𝕜 E := of_fintype_basis ξ + · haveI : FiniteDimensional 𝕜 E := .of_fintype_basis ξ -- first step: thanks to the induction hypothesis, any n-dimensional subspace is equivalent -- to a standard space of dimension n, hence it is complete and therefore closed. have H₁ : ∀ s : Submodule 𝕜 E, finrank 𝕜 s = n → IsClosed (s : Set E) := by @@ -264,7 +264,7 @@ continuous (see `LinearMap.continuous_of_finiteDimensional`), which in turn impl norms are equivalent in finite dimensions. -/ theorem continuous_equivFun_basis [T2Space E] {ι : Type*} [Finite ι] (ξ : Basis ι 𝕜 E) : Continuous ξ.equivFun := - haveI : FiniteDimensional 𝕜 E := of_fintype_basis ξ + haveI : FiniteDimensional 𝕜 E := .of_fintype_basis ξ ξ.equivFun.toLinearMap.continuous_of_finiteDimensional namespace LinearMap diff --git a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean index 1018155f553a7..46a5798a92a07 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean @@ -84,7 +84,7 @@ Hausdorff measure, Hausdorff dimension, dimension open scoped MeasureTheory ENNReal NNReal Topology -open MeasureTheory MeasureTheory.Measure Set TopologicalSpace FiniteDimensional Filter +open MeasureTheory MeasureTheory.Measure Set TopologicalSpace Module Filter variable {ι X Y : Type*} [EMetricSpace X] [EMetricSpace Y] @@ -441,7 +441,7 @@ theorem dimH_univ_pi_fin (n : ℕ) : dimH (univ : Set (Fin n → ℝ)) = n := by theorem dimH_of_mem_nhds {x : E} {s : Set E} (h : s ∈ 𝓝 x) : dimH s = finrank ℝ E := by have e : E ≃L[ℝ] Fin (finrank ℝ E) → ℝ := - ContinuousLinearEquiv.ofFinrankEq (FiniteDimensional.finrank_fin_fun ℝ).symm + ContinuousLinearEquiv.ofFinrankEq (Module.finrank_fin_fun ℝ).symm rw [← e.dimH_image] refine le_antisymm ?_ ?_ · exact (dimH_mono (subset_univ _)).trans_eq (dimH_univ_pi_fin _) @@ -459,7 +459,7 @@ theorem dimH_univ_eq_finrank : dimH (univ : Set E) = finrank ℝ E := dimH_of_mem_nhds (@univ_mem _ (𝓝 0)) theorem dimH_univ : dimH (univ : Set ℝ) = 1 := by - rw [dimH_univ_eq_finrank ℝ, FiniteDimensional.finrank_self, Nat.cast_one] + rw [dimH_univ_eq_finrank ℝ, Module.finrank_self, Nat.cast_one] variable {E} From 7a0b6916909d56dfe75d51505a77692205276f3a Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 3 Oct 2024 20:03:10 +0000 Subject: [PATCH 162/174] chore: three cdots (#17385) --- Mathlib/Data/ENNReal/Real.lean | 4 ++-- Mathlib/Data/Real/ConjExponents.lean | 2 +- Mathlib/Order/Interval/Set/Basic.lean | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index dbe929eb16e42..e9cc940e48612 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -518,8 +518,8 @@ theorem toReal_sSup (s : Set ℝ≥0∞) (hf : ∀ r ∈ s, r ≠ ∞) : obtain rfl | ha := eq_or_ne a ∞ · simp rw [le_iInf_iff, le_ofReal_iff_toReal_le ha, le_ciInf_iff ⟨0, by simpa [mem_lowerBounds]⟩] - exact forall_congr' fun i ↦ (le_ofReal_iff_toReal_le ha (h _)).symm - exact Real.iInf_nonneg h + · exact forall_congr' fun i ↦ (le_ofReal_iff_toReal_le ha (h _)).symm + · exact Real.iInf_nonneg h theorem iInf_add : iInf f + a = ⨅ i, f i + a := le_antisymm (le_iInf fun _ => add_le_add (iInf_le _ _) <| le_rfl) diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index d315cee34de3e..8ade2aebf3b1e 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -255,7 +255,7 @@ variable {a b p q : ℝ≥0∞} (h : p.IsConjExponent q) refine ⟨fun h ↦ ⟨?_, ?_⟩, ?_⟩ · simpa using (ENNReal.lt_add_right (fun hp ↦ by simp [hp] at h) <| by simp).trans_eq h · rw [← coe_inv, ← coe_inv] at h - norm_cast at h + · norm_cast at h all_goals rintro rfl; simp at h · rintro ⟨hp, h⟩ rw [← coe_inv (zero_lt_one.trans hp).ne', ← coe_inv, ← coe_add, h, coe_one] diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 07d4cb76870cc..66feeca1fd5ef 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -1567,7 +1567,7 @@ theorem Ioc_union_Ioc_symm : Ioc a b ∪ Ioc b a = Ioc (min a b) (max a b) := by theorem Ioc_union_Ioc_union_Ioc_cycle : Ioc a b ∪ Ioc b c ∪ Ioc c a = Ioc (min a (min b c)) (max a (max b c)) := by rw [Ioc_union_Ioc, Ioc_union_Ioc] - ac_rfl + · ac_rfl all_goals solve_by_elim (config := { maxDepth := 5 }) [min_le_of_left_le, min_le_of_right_le, le_max_of_le_left, le_max_of_le_right, le_refl] From aeefd81732940f1ef620cba70dbe9ca2ba208fb4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 20:47:19 +0000 Subject: [PATCH 163/174] refactor: Let `positivity` handle `ENNReal`-valued `ofNat` (#17212) The meta code was looking for `StrictOrderedSemiring` instead of the weaker `OrderedSemiring` + `Nontrivial`, which is enough. From LeanAPAP --- Mathlib/Tactic/Positivity/Core.lean | 7 ++++--- test/positivity.lean | 3 +++ 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/Positivity/Core.lean b/Mathlib/Tactic/Positivity/Core.lean index a03631e11fca8..5ea9541b1909c 100644 --- a/Mathlib/Tactic/Positivity/Core.lean +++ b/Mathlib/Tactic/Positivity/Core.lean @@ -122,7 +122,7 @@ variable {A : Type*} {e : A} lemma lt_of_le_of_ne' {a b : A} [PartialOrder A] : (a : A) ≤ b → b ≠ a → a < b := fun h₁ h₂ => lt_of_le_of_ne h₁ h₂.symm -lemma pos_of_isNat {n : ℕ} [StrictOrderedSemiring A] +lemma pos_of_isNat {n : ℕ} [OrderedSemiring A] [Nontrivial A] (h : NormNum.IsNat e n) (w : Nat.ble 1 n = true) : 0 < (e : A) := by rw [NormNum.IsNat.to_eq h rfl] apply Nat.cast_pos.2 @@ -184,11 +184,12 @@ def normNumPositivity (e : Q($α)) : MetaM (Strictness zα pα e) := catchNone d | .isBool .. => failure | .isNat _ lit p => if 0 < lit.natLit! then - let _a ← synthInstanceQ q(StrictOrderedSemiring $α) + let _a ← synthInstanceQ q(OrderedSemiring $α) + let _a ← synthInstanceQ q(Nontrivial $α) assumeInstancesCommute have p : Q(NormNum.IsNat $e $lit) := p haveI' p' : Nat.ble 1 $lit =Q true := ⟨⟩ - pure (.positive q(@pos_of_isNat $α _ _ _ $p $p')) + pure (.positive q(@pos_of_isNat $α _ _ _ _ $p $p')) else let _a ← synthInstanceQ q(OrderedSemiring $α) assumeInstancesCommute diff --git a/test/positivity.lean b/test/positivity.lean index d75fda862bd69..fc773c8abd962 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -23,6 +23,9 @@ example : 0 ≤ 3 := by positivity example : 0 < 3 := by positivity +example : (0 : ℝ≥0∞) < 1 := by positivity +example : (0 : ℝ≥0∞) < 2 := by positivity + /- ## Goals working directly from a hypothesis -/ -- set_option trace.Meta.debug true -- sudo set_option trace.Tactic.positivity true From 3ca1060d7e48260d1b1b7a0733f887e455455226 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 3 Oct 2024 20:47:20 +0000 Subject: [PATCH 164/174] chore: Rename `UniformEmbedding` to `IsUniformEmbedding` (#17295) `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 `DenseEmbedding` 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. --- .../Instances.lean | 24 +-- .../NonUnital.lean | 2 +- .../Restrict.lean | 30 +-- Mathlib/Analysis/CStarAlgebra/Multiplier.lean | 10 +- Mathlib/Analysis/Complex/Basic.lean | 9 +- .../InnerProductSpace/LinearPMap.lean | 2 +- .../Analysis/Normed/Algebra/Unitization.lean | 7 +- .../Normed/Algebra/UnitizationL1.lean | 2 +- .../Normed/Operator/ContinuousLinearMap.lean | 10 +- .../OperatorNorm/Completeness.lean | 2 +- Mathlib/Analysis/Quaternion.lean | 4 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- .../Function/SimpleFuncDenseLp.lean | 10 +- .../Integral/IntegralEqImproper.lean | 4 +- .../MeasureTheory/Integral/SetIntegral.lean | 2 +- .../Algebra/Module/Alternating/Topology.lean | 34 ++-- Mathlib/Topology/Algebra/Module/Basic.lean | 17 +- .../Algebra/Module/FiniteDimension.lean | 6 +- .../Algebra/Module/Multilinear/Topology.lean | 32 +-- .../Algebra/Module/StrongTopology.lean | 38 ++-- .../Topology/Algebra/SeparationQuotient.lean | 5 +- Mathlib/Topology/Algebra/UniformField.lean | 6 +- Mathlib/Topology/Algebra/UniformGroup.lean | 5 +- Mathlib/Topology/ContinuousMap/Bounded.lean | 4 +- Mathlib/Topology/ContinuousMap/Compact.lean | 7 +- .../ContinuousMap/ContinuousMapZero.lean | 32 +-- Mathlib/Topology/EMetricSpace/Basic.lean | 26 ++- Mathlib/Topology/EMetricSpace/Defs.lean | 2 +- Mathlib/Topology/Instances/Int.lean | 9 +- Mathlib/Topology/Instances/Nat.lean | 9 +- Mathlib/Topology/Instances/PNat.lean | 5 +- Mathlib/Topology/Instances/Rat.lean | 27 ++- .../Topology/MetricSpace/Antilipschitz.lean | 9 +- Mathlib/Topology/MetricSpace/Basic.lean | 25 ++- Mathlib/Topology/MetricSpace/Closeds.lean | 13 +- Mathlib/Topology/MetricSpace/Dilation.lean | 10 +- Mathlib/Topology/MetricSpace/Isometry.lean | 15 +- .../Topology/MetricSpace/Pseudo/Basic.lean | 17 +- Mathlib/Topology/UniformSpace/Ascoli.lean | 3 +- Mathlib/Topology/UniformSpace/Basic.lean | 2 +- .../UniformSpace/CompactConvergence.lean | 37 ++-- .../Topology/UniformSpace/CompareReals.lean | 2 +- .../UniformSpace/CompleteSeparated.lean | 7 +- Mathlib/Topology/UniformSpace/Completion.lean | 14 +- Mathlib/Topology/UniformSpace/Equiv.lean | 12 +- .../UniformConvergenceTopology.lean | 15 +- .../UniformSpace/UniformEmbedding.lean | 185 ++++++++++++------ scripts/no_lints_prime_decls.txt | 4 +- 48 files changed, 479 insertions(+), 275 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index d1e7e6f1539c8..a085109299f03 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -218,7 +218,7 @@ lemma QuasispectrumRestricts.isSelfAdjoint (a : A) (ha : QuasispectrumRestricts instance IsSelfAdjoint.instNonUnitalContinuousFunctionalCalculus : NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop) := QuasispectrumRestricts.cfc (q := IsStarNormal) (p := IsSelfAdjoint) Complex.reCLM - Complex.isometry_ofReal.uniformEmbedding (.zero _) + Complex.isometry_ofReal.isUniformEmbedding (.zero _) (fun _ ↦ isSelfAdjoint_iff_isStarNormal_and_quasispectrumRestricts) end SelfAdjointNonUnital @@ -264,7 +264,7 @@ lemma SpectrumRestricts.isSelfAdjoint (a : A) (ha : SpectrumRestricts a Complex. instance IsSelfAdjoint.instContinuousFunctionalCalculus : ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop) := SpectrumRestricts.cfc (q := IsStarNormal) (p := IsSelfAdjoint) Complex.reCLM - Complex.isometry_ofReal.uniformEmbedding (.zero _) + Complex.isometry_ofReal.isUniformEmbedding (.zero _) (fun _ ↦ isSelfAdjoint_iff_isStarNormal_and_spectrumRestricts) lemma IsSelfAdjoint.spectrum_nonempty {A : Type*} [Ring A] [StarRing A] @@ -313,7 +313,7 @@ open NNReal in instance Nonneg.instNonUnitalContinuousFunctionalCalculus : NonUnitalContinuousFunctionalCalculus ℝ≥0 (fun x : A ↦ 0 ≤ x) := QuasispectrumRestricts.cfc (q := IsSelfAdjoint) ContinuousMap.realToNNReal - uniformEmbedding_subtype_val le_rfl + isUniformEmbedding_subtype_val le_rfl (fun _ ↦ nonneg_iff_isSelfAdjoint_and_quasispectrumRestricts) open NNReal in @@ -359,7 +359,7 @@ open NNReal in instance Nonneg.instContinuousFunctionalCalculus : ContinuousFunctionalCalculus ℝ≥0 (fun x : A ↦ 0 ≤ x) := SpectrumRestricts.cfc (q := IsSelfAdjoint) ContinuousMap.realToNNReal - uniformEmbedding_subtype_val le_rfl (fun _ ↦ nonneg_iff_isSelfAdjoint_and_spectrumRestricts) + isUniformEmbedding_subtype_val le_rfl (fun _ ↦ nonneg_iff_isSelfAdjoint_and_spectrumRestricts) end Nonneg @@ -605,14 +605,14 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] lemma cfcHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcHom ha = ha.spectrumRestricts.starAlgHom (cfcHom ha.isStarNormal) (f := Complex.reCLM) := - ha.spectrumRestricts.cfcHom_eq_restrict Complex.isometry_ofReal.uniformEmbedding + ha.spectrumRestricts.cfcHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfc_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : cfc f a = cfc (fun x ↦ f x.re : ℂ → ℂ) a := by replace ha : IsSelfAdjoint a := ha -- hack to avoid issues caused by autoParam exact ha.spectrumRestricts.cfc_eq_restrict (f := Complex.reCLM) - Complex.isometry_ofReal.uniformEmbedding ha ha.isStarNormal f + Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal f end RealEqComplex @@ -626,14 +626,14 @@ variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module lemma cfcₙHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcₙHom ha = (ha.quasispectrumRestricts.2).nonUnitalStarAlgHom (cfcₙHom ha.isStarNormal) (f := Complex.reCLM) := - ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict Complex.isometry_ofReal.uniformEmbedding + ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfcₙ_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : cfcₙ f a = cfcₙ (fun x ↦ f x.re : ℂ → ℂ) a := by replace ha : IsSelfAdjoint a := ha -- hack to avoid issues caused by autoParam exact ha.quasispectrumRestricts.2.cfcₙ_eq_restrict (f := Complex.reCLM) - Complex.isometry_ofReal.uniformEmbedding ha ha.isStarNormal f + Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal f end RealEqComplexNonUnital @@ -650,13 +650,13 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [PartialOrder A] [StarRing A] lemma cfcHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcHom ha = (SpectrumRestricts.nnreal_of_nonneg ha).starAlgHom (cfcHom (IsSelfAdjoint.of_nonneg ha)) := by - apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict uniformEmbedding_subtype_val + apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict isUniformEmbedding_subtype_val lemma cfc_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfc f a = cfc (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam apply (SpectrumRestricts.nnreal_of_nonneg ha).cfc_eq_restrict - uniformEmbedding_subtype_val ha (.of_nonneg ha) + isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqReal @@ -675,13 +675,13 @@ lemma cfcₙHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcₙHom ha = (QuasispectrumRestricts.nnreal_of_nonneg ha).nonUnitalStarAlgHom (cfcₙHom (IsSelfAdjoint.of_nonneg ha)) := by apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙHom_eq_restrict - uniformEmbedding_subtype_val + isUniformEmbedding_subtype_val lemma cfcₙ_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfcₙ f a = cfcₙ (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙ_eq_restrict - uniformEmbedding_subtype_val ha (.of_nonneg ha) + isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqRealNonUnital diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 3ebb272f0dee6..3bcd0795b06d8 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -692,7 +692,7 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : let f : C(spectrum R a, σₙ R a) := ⟨_, continuous_inclusion <| spectrum_subset_quasispectrum R a⟩ refine (cfcHom_closedEmbedding ha).comp <| - (UniformInducing.uniformEmbedding ⟨?_⟩).toClosedEmbedding + (UniformInducing.isUniformEmbedding ⟨?_⟩).toClosedEmbedding 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, diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index 9a2bf73293ea6..cf36028a7922e 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -84,17 +84,17 @@ 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) - (halg : UniformEmbedding (algebraMap R S)) : + (halg : IsUniformEmbedding (algebraMap R S)) : ClosedEmbedding (h.starAlgHom φ) := - hφ.comp <| UniformEmbedding.toClosedEmbedding <| .comp - (ContinuousMap.uniformEmbedding_comp _ halg) - (UniformEquiv.arrowCongr h.homeomorph.symm (.refl _) |>.uniformEmbedding) + hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + (ContinuousMap.isUniformEmbedding_comp _ halg) + (UniformEquiv.arrowCongr h.homeomorph.symm (.refl _) |>.isUniformEmbedding) /-- 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 `ContinuousFunctionalCalculus R p`. -/ -protected theorem cfc (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) (h0 : p 0) +protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h0 : p 0) (h : ∀ a, p a ↔ q a ∧ SpectrumRestricts a f) : ContinuousFunctionalCalculus R p where predicate_zero := h0 @@ -133,14 +133,14 @@ protected theorem cfc (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) ( variable [ContinuousFunctionalCalculus R p] [UniqueContinuousFunctionalCalculus R A] -lemma cfcHom_eq_restrict (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) +lemma cfcHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {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.starAlgHom_id (cfcHom_id hqa) -lemma cfc_eq_restrict (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) {a : A} (hpa : p a) +lemma cfc_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} (hpa : p a) (hqa : q a) (h : SpectrumRestricts a f) (g : R → R) : cfc g a = cfc (fun x ↦ algebraMap R S (g (f x))) a := by by_cases hg : ContinuousOn g (spectrum R a) @@ -218,12 +218,12 @@ variable [CompleteSpace R] lemma closedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} (hφ : ClosedEmbedding φ) {f : C(S, R)} (h : QuasispectrumRestricts a f) - (halg : UniformEmbedding (algebraMap R S)) : + (halg : IsUniformEmbedding (algebraMap R S)) : ClosedEmbedding (h.nonUnitalStarAlgHom φ) := by have : h.homeomorph.symm 0 = 0 := Subtype.ext (map_zero <| algebraMap _ _) - refine hφ.comp <| UniformEmbedding.toClosedEmbedding <| .comp - (ContinuousMapZero.uniformEmbedding_comp _ halg) - (UniformEquiv.arrowCongrLeft₀ h.homeomorph.symm this |>.uniformEmbedding) + refine hφ.comp <| IsUniformEmbedding.toClosedEmbedding <| .comp + (ContinuousMapZero.isUniformEmbedding_comp _ halg) + (UniformEquiv.arrowCongrLeft₀ h.homeomorph.symm this |>.isUniformEmbedding) variable [IsScalarTower R A A] [SMulCommClass R A A] @@ -231,7 +231,7 @@ variable [IsScalarTower R A A] [SMulCommClass R A A] characterized by: `q a` and the quasispectrum of `a` restricts to the scalar subring `R` via `f : C(S, R)`, then we can get a restricted functional calculus `NonUnitalContinuousFunctionalCalculus R p`. -/ -protected theorem cfc (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) (h0 : p 0) +protected theorem cfc (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) (h0 : p 0) (h : ∀ a, p a ↔ q a ∧ QuasispectrumRestricts a f) : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := h0 @@ -275,15 +275,15 @@ protected theorem cfc (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) ( variable [NonUnitalContinuousFunctionalCalculus R p] variable [UniqueNonUnitalContinuousFunctionalCalculus R A] -lemma cfcₙHom_eq_restrict (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) {a : A} +lemma cfcₙHom_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} (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.nonUnitalStarAlgHom_id (cfcₙHom_id hqa) -lemma cfcₙ_eq_restrict (f : C(S, R)) (halg : UniformEmbedding (algebraMap R S)) {a : A} (hpa : p a) - (hqa : q a) (h : QuasispectrumRestricts a f) (g : R → R) : +lemma cfcₙ_eq_restrict (f : C(S, R)) (halg : IsUniformEmbedding (algebraMap R S)) {a : A} + (hpa : p a) (hqa : q a) (h : QuasispectrumRestricts a f) (g : R → R) : cfcₙ g a = cfcₙ (fun x ↦ algebraMap R S (g (f x))) a := by by_cases hg : ContinuousOn g (σₙ R a) ∧ g 0 = 0 · obtain ⟨hg, hg0⟩ := hg diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index 248d3fa9e31ac..d9c622e26c2c4 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -534,11 +534,15 @@ instance instNormedSpace : NormedSpace 𝕜 𝓜(𝕜, A) := instance instNormedAlgebra : NormedAlgebra 𝕜 𝓜(𝕜, A) := { DoubleCentralizer.instAlgebra, DoubleCentralizer.instNormedSpace with } -theorem uniformEmbedding_toProdMulOpposite : UniformEmbedding (@toProdMulOpposite 𝕜 A _ _ _ _ _) := - uniformEmbedding_comap toProdMulOpposite_injective +theorem isUniformEmbedding_toProdMulOpposite : + IsUniformEmbedding (toProdMulOpposite (𝕜 := 𝕜) (A := A)) := + isUniformEmbedding_comap toProdMulOpposite_injective + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toProdMulOpposite := isUniformEmbedding_toProdMulOpposite instance [CompleteSpace A] : CompleteSpace 𝓜(𝕜, A) := by - rw [completeSpace_iff_isComplete_range uniformEmbedding_toProdMulOpposite.toUniformInducing] + rw [completeSpace_iff_isComplete_range isUniformEmbedding_toProdMulOpposite.toUniformInducing] apply IsClosed.isComplete simp only [range_toProdMulOpposite, Set.setOf_forall] refine isClosed_iInter fun x => isClosed_iInter fun y => isClosed_eq ?_ ?_ diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 28550b2fd6baa..b2b77580329f5 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -208,11 +208,14 @@ theorem antilipschitz_equivRealProd : AntilipschitzWith (NNReal.sqrt 2) equivRea AddMonoidHomClass.antilipschitz_of_bound equivRealProdLm fun z ↦ by simpa only [Real.coe_sqrt, NNReal.coe_ofNat] using abs_le_sqrt_two_mul_max z -theorem uniformEmbedding_equivRealProd : UniformEmbedding equivRealProd := - antilipschitz_equivRealProd.uniformEmbedding lipschitz_equivRealProd.uniformContinuous +theorem isUniformEmbedding_equivRealProd : IsUniformEmbedding equivRealProd := + antilipschitz_equivRealProd.isUniformEmbedding lipschitz_equivRealProd.uniformContinuous + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_equivRealProd := isUniformEmbedding_equivRealProd instance : CompleteSpace ℂ := - (completeSpace_congr uniformEmbedding_equivRealProd).mpr inferInstance + (completeSpace_congr isUniformEmbedding_equivRealProd).mpr inferInstance instance instT2Space : T2Space ℂ := TopologicalSpace.t2Space_of_metrizableSpace diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 4fa2bc6e39428..4ab4ac3e7b9f6 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -103,7 +103,7 @@ variable (hT : Dense (T.domain : Set E)) /-- The unique continuous extension of the operator `adjointDomainMkCLM` to `E`. -/ def adjointDomainMkCLMExtend (y : T.adjointDomain) : E →L[𝕜] 𝕜 := (T.adjointDomainMkCLM y).extend (Submodule.subtypeL T.domain) hT.denseRange_val - uniformEmbedding_subtype_val.toUniformInducing + isUniformEmbedding_subtype_val.toUniformInducing @[simp] theorem adjointDomainMkCLMExtend_apply (y : T.adjointDomain) (x : T.domain) : diff --git a/Mathlib/Analysis/Normed/Algebra/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean index 8b5d489952f9a..19b11e9819409 100644 --- a/Mathlib/Analysis/Normed/Algebra/Unitization.lean +++ b/Mathlib/Analysis/Normed/Algebra/Unitization.lean @@ -208,11 +208,14 @@ def uniformEquivProd : (Unitization 𝕜 A) ≃ᵤ (𝕜 × A) := instance instBornology : Bornology (Unitization 𝕜 A) := Bornology.induced <| addEquiv 𝕜 A -theorem uniformEmbedding_addEquiv {𝕜} [NontriviallyNormedField 𝕜] : - UniformEmbedding (addEquiv 𝕜 A) where +theorem isUniformEmbedding_addEquiv {𝕜} [NontriviallyNormedField 𝕜] : + IsUniformEmbedding (addEquiv 𝕜 A) where comap_uniformity := rfl inj := (addEquiv 𝕜 A).injective +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_addEquiv := isUniformEmbedding_addEquiv + /-- `Unitization 𝕜 A` is complete whenever `𝕜` and `A` are also. -/ instance instCompleteSpace [CompleteSpace 𝕜] [CompleteSpace A] : CompleteSpace (Unitization 𝕜 A) := diff --git a/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean index 4a0ab143ce2a1..bb8f42626297c 100644 --- a/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean +++ b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean @@ -50,7 +50,7 @@ noncomputable def uniformEquiv_unitization_addEquiv_prod : instance instCompleteSpace [CompleteSpace 𝕜] [CompleteSpace A] : CompleteSpace (WithLp 1 (Unitization 𝕜 A)) := - completeSpace_congr (uniformEquiv_unitization_addEquiv_prod 𝕜 A).uniformEmbedding |>.mpr + completeSpace_congr (uniformEquiv_unitization_addEquiv_prod 𝕜 A).isUniformEmbedding |>.mpr CompleteSpace.prod variable {𝕜 A} diff --git a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean index eedbaba72b905..64c5ba0eb60b5 100644 --- a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean +++ b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean @@ -153,9 +153,13 @@ variable [Ring 𝕜] [Ring 𝕜₂] variable [NormedAddCommGroup E] [NormedAddCommGroup F] [Module 𝕜 E] [Module 𝕜₂ F] variable {σ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ] F) (x y z : E) -theorem ContinuousLinearMap.uniformEmbedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : - UniformEmbedding f := - (AddMonoidHomClass.antilipschitz_of_bound f hf).uniformEmbedding f.uniformContinuous +theorem ContinuousLinearMap.isUniformEmbedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : + IsUniformEmbedding f := + (AddMonoidHomClass.antilipschitz_of_bound f hf).isUniformEmbedding f.uniformContinuous + +@[deprecated (since := "2024-10-01")] +alias ContinuousLinearMap.uniformEmbedding_of_bound := + ContinuousLinearMap.isUniformEmbedding_of_bound end Normed diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 716b0f8e64316..9c9517852f03e 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -230,7 +230,7 @@ variable {N : ℝ≥0} (h_e : ∀ x, ‖x‖ ≤ N * ‖e x‖) [RingHomIsometri /-- If a dense embedding `e : E →L[𝕜] G` expands the norm by a constant factor `N⁻¹`, then the norm of the extension of `f` along `e` is bounded by `N * ‖f‖`. -/ theorem opNorm_extend_le : - ‖f.extend e h_dense (uniformEmbedding_of_bound _ h_e).toUniformInducing‖ ≤ N * ‖f‖ := by + ‖f.extend e h_dense (isUniformEmbedding_of_bound _ h_e).toUniformInducing‖ ≤ N * ‖f‖ := by -- Add `opNorm_le_of_dense`? refine opNorm_le_bound _ ?_ (isClosed_property h_dense (isClosed_le ?_ ?_) fun x ↦ ?_) · cases le_total 0 N with diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index e342a1859bb60..3466df2251849 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -195,8 +195,8 @@ theorem continuous_im : Continuous fun q : ℍ => q.im := by simpa only [← sub_self_re] using continuous_id.sub (continuous_coe.comp continuous_re) instance : CompleteSpace ℍ := - haveI : UniformEmbedding linearIsometryEquivTuple.toLinearEquiv.toEquiv.symm := - linearIsometryEquivTuple.toContinuousLinearEquiv.symm.uniformEmbedding + haveI : IsUniformEmbedding linearIsometryEquivTuple.toLinearEquiv.toEquiv.symm := + linearIsometryEquivTuple.toContinuousLinearEquiv.symm.isUniformEmbedding (completeSpace_congr this).1 inferInstance section infinite_sum diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index ed9015bf0ac57..5410247494af6 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -1064,7 +1064,7 @@ theorem MeasureTheory.Memℒp.of_comp_antilipschitzWith {α E F} {K'} [Measurabl rw [← dist_zero_right, ← dist_zero_right, ← g0] apply hg'.le_mul_dist have B : AEStronglyMeasurable f μ := - (hg'.uniformEmbedding hg).embedding.aestronglyMeasurable_comp_iff.1 hL.1 + (hg'.isUniformEmbedding hg).embedding.aestronglyMeasurable_comp_iff.1 hL.1 exact hL.of_le_mul B (Filter.Eventually.of_forall A) namespace LipschitzWith diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 29f9a7d7b2d9b..72dd2f1509641 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -679,16 +679,18 @@ variable [Fact (1 ≤ p)] protected theorem uniformContinuous : UniformContinuous ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := uniformContinuous_comap -protected theorem uniformEmbedding : UniformEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - uniformEmbedding_comap Subtype.val_injective +lemma isUniformEmbedding : IsUniformEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := + isUniformEmbedding_comap Subtype.val_injective + +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding protected theorem uniformInducing : UniformInducing ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := - simpleFunc.uniformEmbedding.toUniformInducing + simpleFunc.isUniformEmbedding.toUniformInducing lemma isDenseEmbedding (hp_ne_top : p ≠ ∞) : IsDenseEmbedding ((↑) : Lp.simpleFunc E p μ → Lp E p μ) := by borelize E - apply simpleFunc.uniformEmbedding.isDenseEmbedding + apply simpleFunc.isUniformEmbedding.isDenseEmbedding intro f rw [mem_closure_iff_seq_limit] have hfi' : Memℒp f p μ := Lp.memℒp f diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index 902812d6d999d..ac2ffb0b0ce8f 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -712,7 +712,7 @@ theorem tendsto_zero_of_hasDerivAt_of_integrableOn_Ioi rw [← top_le_iff, ← volume_Ici (a := b)] exact measure_mono hb rwa [B, ← Embedding.tendsto_nhds_iff] at A - exact (Completion.uniformEmbedding_coe E).embedding + exact (Completion.isUniformEmbedding_coe E).embedding variable [CompleteSpace E] @@ -909,7 +909,7 @@ theorem tendsto_zero_of_hasDerivAt_of_integrableOn_Iic rw [← volume_Iic (a := b)] exact measure_mono hb rwa [B, ← Embedding.tendsto_nhds_iff] at A - exact (Completion.uniformEmbedding_coe E).embedding + exact (Completion.isUniformEmbedding_coe E).embedding variable [CompleteSpace E] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 4af2e6de396cb..be099131dfc9b 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1260,7 +1260,7 @@ variable [NormedSpace ℝ F] [NormedSpace ℝ E] theorem integral_comp_comm (L : E ≃L[𝕜] F) (φ : X → E) : ∫ x, L (φ x) ∂μ = L (∫ x, φ x ∂μ) := by have : CompleteSpace E ↔ CompleteSpace F := - completeSpace_congr (e := L.toEquiv) L.uniformEmbedding + completeSpace_congr (e := L.toEquiv) L.isUniformEmbedding obtain ⟨_, _⟩|⟨_, _⟩ := iff_iff_and_or_not_and_not.mp this · exact L.toContinuousLinearMap.integral_comp_comm' L.antilipschitz _ · simp [integral, *] diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean index 747f1739055fe..9f613d081db1f 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Topology.lean @@ -47,14 +47,17 @@ variable [UniformSpace F] [UniformAddGroup F] instance instUniformSpace : UniformSpace (E [⋀^ι]→L[𝕜] F) := .comap toContinuousMultilinearMap inferInstance -lemma uniformEmbedding_toContinuousMultilinearMap : - UniformEmbedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → _) where +lemma isUniformEmbedding_toContinuousMultilinearMap : + IsUniformEmbedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → _) where inj := toContinuousMultilinearMap_injective comap_uniformity := rfl +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toContinuousMultilinearMap := isUniformEmbedding_toContinuousMultilinearMap + lemma uniformContinuous_toContinuousMultilinearMap : UniformContinuous (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F) → _) := - uniformEmbedding_toContinuousMultilinearMap.uniformContinuous + isUniformEmbedding_toContinuousMultilinearMap.uniformContinuous theorem uniformContinuous_coe_fun [ContinuousSMul 𝕜 E] : UniformContinuous (DFunLike.coe : (E [⋀^ι]→L[𝕜] F) → (ι → E) → F) := @@ -66,13 +69,13 @@ theorem uniformContinuous_eval_const [ContinuousSMul 𝕜 E] (x : ι → E) : uniformContinuous_pi.1 uniformContinuous_coe_fun x instance instUniformAddGroup : UniformAddGroup (E [⋀^ι]→L[𝕜] F) := - uniformEmbedding_toContinuousMultilinearMap.uniformAddGroup + isUniformEmbedding_toContinuousMultilinearMap.uniformAddGroup (toContinuousMultilinearMapLinear (R := ℕ)) instance instUniformContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜 M F] [ContinuousConstSMul M F] : UniformContinuousConstSMul M (E [⋀^ι]→L[𝕜] F) := - uniformEmbedding_toContinuousMultilinearMap.uniformContinuousConstSMul fun _ _ ↦ rfl + isUniformEmbedding_toContinuousMultilinearMap.uniformContinuousConstSMul fun _ _ ↦ rfl section CompleteSpace @@ -83,7 +86,7 @@ theorem completeSpace (h : RestrictGenTopology {s : Set (ι → E) | IsVonNBound CompleteSpace (E [⋀^ι]→L[𝕜] F) := by have := ContinuousMultilinearMap.completeSpace (F := F) h rw [completeSpace_iff_isComplete_range - uniformEmbedding_toContinuousMultilinearMap.toUniformInducing] + isUniformEmbedding_toContinuousMultilinearMap.toUniformInducing] apply isClosed_range_toContinuousMultilinearMap.isComplete instance instCompleteSpace [TopologicalAddGroup E] [SequentialSpace (ι → E)] : @@ -97,15 +100,18 @@ section RestrictScalars variable (𝕜' : Type*) [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜' 𝕜] [Module 𝕜' E] [IsScalarTower 𝕜' 𝕜 E] [Module 𝕜' F] [IsScalarTower 𝕜' 𝕜 F] [ContinuousSMul 𝕜 E] -theorem uniformEmbedding_restrictScalars : - UniformEmbedding (restrictScalars 𝕜' : E [⋀^ι]→L[𝕜] F → E [⋀^ι]→L[𝕜'] F) := by - rw [← uniformEmbedding_toContinuousMultilinearMap.of_comp_iff] - exact (ContinuousMultilinearMap.uniformEmbedding_restrictScalars 𝕜').comp - uniformEmbedding_toContinuousMultilinearMap +theorem isUniformEmbedding_restrictScalars : + IsUniformEmbedding (restrictScalars 𝕜' : E [⋀^ι]→L[𝕜] F → E [⋀^ι]→L[𝕜'] F) := by + rw [← isUniformEmbedding_toContinuousMultilinearMap.of_comp_iff] + exact (ContinuousMultilinearMap.isUniformEmbedding_restrictScalars 𝕜').comp + isUniformEmbedding_toContinuousMultilinearMap + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_restrictScalars := isUniformEmbedding_restrictScalars theorem uniformContinuous_restrictScalars : UniformContinuous (restrictScalars 𝕜' : E [⋀^ι]→L[𝕜] F → E [⋀^ι]→L[𝕜'] F) := - (uniformEmbedding_restrictScalars 𝕜').uniformContinuous + (isUniformEmbedding_restrictScalars 𝕜').uniformContinuous end RestrictScalars @@ -117,7 +123,7 @@ lemma embedding_toContinuousMultilinearMap : Embedding (toContinuousMultilinearMap : (E [⋀^ι]→L[𝕜] F → _)) := letI := TopologicalAddGroup.toUniformSpace F haveI := comm_topologicalAddGroup_is_uniform (G := F) - uniformEmbedding_toContinuousMultilinearMap.embedding + isUniformEmbedding_toContinuousMultilinearMap.embedding @[continuity, fun_prop] lemma continuous_toContinuousMultilinearMap : @@ -179,7 +185,7 @@ theorem embedding_restrictScalars : Embedding (restrictScalars 𝕜' : E [⋀^ι]→L[𝕜] F → E [⋀^ι]→L[𝕜'] F) := letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform - (uniformEmbedding_restrictScalars _).embedding + (isUniformEmbedding_restrictScalars _).embedding @[continuity, fun_prop] theorem continuous_restrictScalars : diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 99adb17d2c59c..76d6e898d399b 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1942,22 +1942,27 @@ protected theorem preimage_symm_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Se e ⁻¹' (e.symm ⁻¹' s) = s := e.symm.symm_preimage_preimage s -protected theorem uniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] +lemma isUniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] - [UniformAddGroup E₂] (e : E₁ ≃SL[σ₁₂] E₂) : UniformEmbedding e := - e.toLinearEquiv.toEquiv.uniformEmbedding e.toContinuousLinearMap.uniformContinuous + [UniformAddGroup E₂] (e : E₁ ≃SL[σ₁₂] E₂) : IsUniformEmbedding e := + e.toLinearEquiv.toEquiv.isUniformEmbedding e.toContinuousLinearMap.uniformContinuous e.symm.toContinuousLinearMap.uniformContinuous -protected theorem _root_.LinearEquiv.uniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding + +protected theorem _root_.LinearEquiv.isUniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] [UniformAddGroup E₂] (e : E₁ ≃ₛₗ[σ₁₂] E₂) - (h₁ : Continuous e) (h₂ : Continuous e.symm) : UniformEmbedding e := - ContinuousLinearEquiv.uniformEmbedding + (h₁ : Continuous e) (h₂ : Continuous e.symm) : IsUniformEmbedding e := + ContinuousLinearEquiv.isUniformEmbedding ({ e with continuous_toFun := h₁ continuous_invFun := h₂ } : E₁ ≃SL[σ₁₂] E₂) +@[deprecated (since := "2024-10-01")] +alias _root_.LinearEquiv.uniformEmbedding := _root_.LinearEquiv.isUniformEmbedding + /-- Create a `ContinuousLinearEquiv` from two `ContinuousLinearMap`s that are inverse of each other. -/ def equivOfInverse (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 3c7e1401aa012..2cf817161d306 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -204,13 +204,13 @@ private theorem continuous_equivFun_basis_aux [T2Space E] {ι : Type v} [Fintype intro s s_dim letI : UniformAddGroup s := s.toAddSubgroup.uniformAddGroup let b := Basis.ofVectorSpace 𝕜 s - have U : UniformEmbedding b.equivFun.symm.toEquiv := by + have U : IsUniformEmbedding b.equivFun.symm.toEquiv := by have : Fintype.card (Basis.ofVectorSpaceIndex 𝕜 s) = n := by rw [← s_dim] exact (finrank_eq_card_basis b).symm have : Continuous b.equivFun := IH b this exact - b.equivFun.symm.uniformEmbedding b.equivFun.symm.toLinearMap.continuous_on_pi this + b.equivFun.symm.isUniformEmbedding b.equivFun.symm.toLinearMap.continuous_on_pi this have : IsComplete (s : Set E) := completeSpace_coe_iff_isComplete.1 ((completeSpace_congr U).1 inferInstance) exact this.isClosed @@ -490,7 +490,7 @@ variable (𝕜 E : Type*) [NontriviallyNormedField 𝕜] include 𝕜 in theorem FiniteDimensional.complete [FiniteDimensional 𝕜 E] : CompleteSpace E := by set e := ContinuousLinearEquiv.ofFinrankEq (@finrank_fin_fun 𝕜 _ _ (finrank 𝕜 E)).symm - have : UniformEmbedding e.toEquiv.symm := e.symm.uniformEmbedding + have : IsUniformEmbedding e.toEquiv.symm := e.symm.isUniformEmbedding exact (completeSpace_congr this).1 inferInstance variable {𝕜 E} diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean index a403ef7dca6bd..c73eb8c4be001 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Topology.lean @@ -65,18 +65,21 @@ section UniformAddGroup variable [UniformSpace F] [UniformAddGroup F] -lemma uniformEmbedding_toUniformOnFun : - UniformEmbedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) where +lemma isUniformEmbedding_toUniformOnFun : + IsUniformEmbedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) where inj := DFunLike.coe_injective comap_uniformity := rfl +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toUniformOnFun := isUniformEmbedding_toUniformOnFun + lemma embedding_toUniformOnFun : Embedding (toUniformOnFun : ContinuousMultilinearMap 𝕜 E F → _) := - uniformEmbedding_toUniformOnFun.embedding + isUniformEmbedding_toUniformOnFun.embedding theorem uniformContinuous_coe_fun [∀ i, ContinuousSMul 𝕜 (E i)] : UniformContinuous (DFunLike.coe : ContinuousMultilinearMap 𝕜 E F → (Π i, E i) → F) := (UniformOnFun.uniformContinuous_toFun isVonNBounded_covers).comp - uniformEmbedding_toUniformOnFun.uniformContinuous + isUniformEmbedding_toUniformOnFun.uniformContinuous theorem uniformContinuous_eval_const [∀ i, ContinuousSMul 𝕜 (E i)] (x : Π i, E i) : UniformContinuous fun f : ContinuousMultilinearMap 𝕜 E F ↦ f x := @@ -85,13 +88,13 @@ theorem uniformContinuous_eval_const [∀ i, ContinuousSMul 𝕜 (E i)] (x : Π instance instUniformAddGroup : UniformAddGroup (ContinuousMultilinearMap 𝕜 E F) := let φ : ContinuousMultilinearMap 𝕜 E F →+ (Π i, E i) →ᵤ[{s | IsVonNBounded 𝕜 s}] F := { toFun := toUniformOnFun, map_add' := fun _ _ ↦ rfl, map_zero' := rfl } - uniformEmbedding_toUniformOnFun.uniformAddGroup φ + isUniformEmbedding_toUniformOnFun.uniformAddGroup φ instance instUniformContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜 M F] [ContinuousConstSMul M F] : UniformContinuousConstSMul M (ContinuousMultilinearMap 𝕜 E F) := haveI := uniformContinuousConstSMul_of_continuousConstSMul M F - uniformEmbedding_toUniformOnFun.uniformContinuousConstSMul fun _ _ ↦ rfl + isUniformEmbedding_toUniformOnFun.uniformContinuousConstSMul fun _ _ ↦ rfl section CompleteSpace @@ -104,7 +107,7 @@ theorem completeSpace (h : RestrictGenTopology {s : Set (Π i, E i) | IsVonNBoun 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 uniformEmbedding_toUniformOnFun.toUniformInducing, + rw [completeSpace_iff_isComplete_range isUniformEmbedding_toUniformOnFun.toUniformInducing, range_toUniformOnFun] simp only [setOf_and, setOf_forall] apply_rules [IsClosed.isComplete, IsClosed.inter] @@ -126,19 +129,22 @@ variable (𝕜' : Type*) [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜' [∀ i, Module 𝕜' (E i)] [∀ i, IsScalarTower 𝕜' 𝕜 (E i)] [Module 𝕜' F] [IsScalarTower 𝕜' 𝕜 F] [∀ i, ContinuousSMul 𝕜 (E i)] -theorem uniformEmbedding_restrictScalars : - UniformEmbedding +theorem isUniformEmbedding_restrictScalars : + IsUniformEmbedding (restrictScalars 𝕜' : ContinuousMultilinearMap 𝕜 E F → ContinuousMultilinearMap 𝕜' E F) := by letI : NontriviallyNormedField 𝕜 := ⟨let ⟨x, hx⟩ := @NontriviallyNormedField.non_trivial 𝕜' _; ⟨algebraMap 𝕜' 𝕜 x, by simpa⟩⟩ - rw [← uniformEmbedding_toUniformOnFun.of_comp_iff] - convert uniformEmbedding_toUniformOnFun using 4 with s + rw [← isUniformEmbedding_toUniformOnFun.of_comp_iff] + convert isUniformEmbedding_toUniformOnFun using 4 with s exact ⟨fun h ↦ h.extend_scalars _, fun h ↦ h.restrict_scalars _⟩ +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_restrictScalars := isUniformEmbedding_restrictScalars + theorem uniformContinuous_restrictScalars : UniformContinuous (restrictScalars 𝕜' : ContinuousMultilinearMap 𝕜 E F → ContinuousMultilinearMap 𝕜' E F) := - (uniformEmbedding_restrictScalars 𝕜').uniformContinuous + (isUniformEmbedding_restrictScalars 𝕜').uniformContinuous end RestrictScalars @@ -207,7 +213,7 @@ theorem embedding_restrictScalars : (restrictScalars 𝕜' : ContinuousMultilinearMap 𝕜 E F → ContinuousMultilinearMap 𝕜' E F) := letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform - (uniformEmbedding_restrictScalars _).embedding + (isUniformEmbedding_restrictScalars _).embedding @[continuity, fun_prop] theorem continuous_restrictScalars : diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 437453ea31e0d..6bb9968ff157e 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -117,14 +117,17 @@ theorem uniformity_toTopologicalSpace_eq [UniformSpace F] [UniformAddGroup F] ( UniformConvergenceCLM.instTopologicalSpace σ F 𝔖 := rfl -theorem uniformEmbedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : - UniformEmbedding (α := UniformConvergenceCLM σ F 𝔖) (β := E →ᵤ[𝔖] F) DFunLike.coe := +theorem isUniformEmbedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : + IsUniformEmbedding (α := UniformConvergenceCLM σ F 𝔖) (β := E →ᵤ[𝔖] F) DFunLike.coe := ⟨⟨rfl⟩, DFunLike.coe_injective⟩ +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_coeFn := isUniformEmbedding_coeFn + theorem embedding_coeFn [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (Set E)) : Embedding (X := UniformConvergenceCLM σ F 𝔖) (Y := E →ᵤ[𝔖] F) (UniformOnFun.ofFun 𝔖 ∘ DFunLike.coe) := - UniformEmbedding.embedding (uniformEmbedding_coeFn _ _ _) + IsUniformEmbedding.embedding (isUniformEmbedding_coeFn _ _ _) instance instAddCommGroup [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : AddCommGroup (UniformConvergenceCLM σ F 𝔖) := ContinuousLinearMap.addCommGroup @@ -133,7 +136,7 @@ instance instUniformAddGroup [UniformSpace F] [UniformAddGroup F] (𝔖 : Set (S UniformAddGroup (UniformConvergenceCLM σ F 𝔖) := by let φ : (UniformConvergenceCLM σ F 𝔖) →+ E →ᵤ[𝔖] F := ⟨⟨(DFunLike.coe : (UniformConvergenceCLM σ F 𝔖) → E →ᵤ[𝔖] F), rfl⟩, fun _ _ => rfl⟩ - exact (uniformEmbedding_coeFn _ _ _).uniformAddGroup φ + exact (isUniformEmbedding_coeFn _ _ _).uniformAddGroup φ instance instTopologicalAddGroup [TopologicalSpace F] [TopologicalAddGroup F] (𝔖 : Set (Set E)) : TopologicalAddGroup (UniformConvergenceCLM σ F 𝔖) := by @@ -189,7 +192,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 𝔖) := - (uniformEmbedding_coeFn σ F 𝔖).toUniformInducing.uniformContinuousConstSMul fun _ _ ↦ by rfl + (isUniformEmbedding_coeFn σ F 𝔖).toUniformInducing.uniformContinuousConstSMul fun _ _ ↦ by rfl instance instContinuousConstSMul (M : Type*) [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] @@ -276,9 +279,13 @@ protected theorem hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGroup F fun SV => { f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2 } := ContinuousLinearMap.hasBasis_nhds_zero_of_basis (𝓝 0).basis_sets -theorem uniformEmbedding_toUniformOnFun [UniformSpace F] [UniformAddGroup F] : - UniformEmbedding fun f : E →SL[σ] F ↦ UniformOnFun.ofFun {s | Bornology.IsVonNBounded 𝕜₁ s} f := - UniformConvergenceCLM.uniformEmbedding_coeFn .. +theorem isUniformEmbedding_toUniformOnFun [UniformSpace F] [UniformAddGroup F] : + IsUniformEmbedding + fun f : E →SL[σ] F ↦ UniformOnFun.ofFun {s | Bornology.IsVonNBounded 𝕜₁ s} f := + UniformConvergenceCLM.isUniformEmbedding_coeFn .. + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toUniformOnFun := isUniformEmbedding_toUniformOnFun instance uniformContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] @@ -362,15 +369,18 @@ variable [UniformSpace F] [UniformAddGroup F] [Module 𝕜 F] (𝕜' : Type*) [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜' 𝕜] [Module 𝕜' E] [IsScalarTower 𝕜' 𝕜 E] [Module 𝕜' F] [IsScalarTower 𝕜' 𝕜 F] -theorem uniformEmbedding_restrictScalars : - UniformEmbedding (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := by - rw [← uniformEmbedding_toUniformOnFun.of_comp_iff] - convert uniformEmbedding_toUniformOnFun using 4 with s +theorem isUniformEmbedding_restrictScalars : + IsUniformEmbedding (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := by + rw [← isUniformEmbedding_toUniformOnFun.of_comp_iff] + convert isUniformEmbedding_toUniformOnFun using 4 with s exact ⟨fun h ↦ h.extend_scalars _, fun h ↦ h.restrict_scalars _⟩ +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_restrictScalars := isUniformEmbedding_restrictScalars + theorem uniformContinuous_restrictScalars : UniformContinuous (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := - (uniformEmbedding_restrictScalars 𝕜').uniformContinuous + (isUniformEmbedding_restrictScalars 𝕜').uniformContinuous end UniformSpace @@ -382,7 +392,7 @@ theorem embedding_restrictScalars : Embedding (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform - (uniformEmbedding_restrictScalars _).embedding + (isUniformEmbedding_restrictScalars _).embedding @[continuity, fun_prop] theorem continuous_restrictScalars : diff --git a/Mathlib/Topology/Algebra/SeparationQuotient.lean b/Mathlib/Topology/Algebra/SeparationQuotient.lean index 19e68ba576e4f..070cdf516d506 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient.lean @@ -435,10 +435,13 @@ theorem outCLM_uniformInducing : UniformInducing (outCLM K E) := by rw [← uniformInducing_mk.uniformInducing_comp_iff, mk_comp_outCLM] exact uniformInducing_id -theorem outCLM_uniformEmbedding : UniformEmbedding (outCLM K E) where +theorem outCLM_isUniformEmbedding : IsUniformEmbedding (outCLM K E) where inj := outCLM_injective K E toUniformInducing := outCLM_uniformInducing K E +@[deprecated (since := "2024-10-01")] +alias outCLM_uniformEmbedding := outCLM_isUniformEmbedding + theorem outCLM_uniformContinuous : UniformContinuous (outCLM K E) := (outCLM_uniformInducing K E).uniformContinuous diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 27cc2841c5eb2..6746d27d74d8b 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -55,7 +55,7 @@ namespace UniformSpace namespace Completion instance (priority := 100) [T0Space K] : Nontrivial (hat K) := - ⟨⟨0, 1, fun h => zero_ne_one <| (uniformEmbedding_coe K).inj h⟩⟩ + ⟨⟨0, 1, fun h => zero_ne_one <| (isUniformEmbedding_coe K).inj h⟩⟩ variable {K} @@ -145,7 +145,7 @@ theorem mul_hatInv_cancel {x : hat K} (x_ne : x ≠ 0) : x * hatInv x = 1 := by rwa [closure_singleton, mem_singleton_iff] at fxclo instance instField : Field (hat K) where - exists_pair_ne := ⟨0, 1, fun h => zero_ne_one ((uniformEmbedding_coe K).inj h)⟩ + exists_pair_ne := ⟨0, 1, fun h => zero_ne_one ((isUniformEmbedding_coe K).inj h)⟩ mul_inv_cancel := fun x x_ne => by simp only [Inv.inv, if_neg x_ne, mul_hatInv_cancel x_ne] inv_zero := by simp only [Inv.inv, ite_true] -- TODO: use a better defeq @@ -176,7 +176,7 @@ variable (L : Type*) [Field L] [UniformSpace L] [CompletableTopField L] instance Subfield.completableTopField (K : Subfield L) : CompletableTopField K where nice F F_cau inf_F := by let i : K →+* L := K.subtype - have hi : UniformInducing i := uniformEmbedding_subtype_val.toUniformInducing + have hi : UniformInducing i := isUniformEmbedding_subtype_val.toUniformInducing rw [← hi.cauchy_map_iff] at F_cau ⊢ rw [map_comm (show (i ∘ fun x => x⁻¹) = (fun x => x⁻¹) ∘ i by ext; rfl)] apply CompletableTopField.nice _ F_cau diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index 21f75fd7a2065..daf8079c33a30 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -179,13 +179,16 @@ theorem uniformity_translate_mul (a : α) : ((𝓤 α).map fun x : α × α => ( ) @[to_additive] -theorem uniformEmbedding_translate_mul (a : α) : UniformEmbedding fun x : α => x * a := +theorem isUniformEmbedding_translate_mul (a : α) : IsUniformEmbedding fun x : α => x * a := { comap_uniformity := by nth_rw 1 [← uniformity_translate_mul a, comap_map] rintro ⟨p₁, p₂⟩ ⟨q₁, q₂⟩ simp only [Prod.mk.injEq, mul_left_inj, imp_self] inj := mul_left_injective a } +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_translate_mul := isUniformEmbedding_translate_mul + namespace MulOpposite @[to_additive] diff --git a/Mathlib/Topology/ContinuousMap/Bounded.lean b/Mathlib/Topology/ContinuousMap/Bounded.lean index 205125c91bfb4..a1fa188b5cc16 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded.lean @@ -246,7 +246,7 @@ theorem inducing_coeFn : Inducing (UniformFun.ofFun ∘ (⇑) : (α →ᵇ β) UniformFun.tendsto_iff_tendstoUniformly] simp [comp_def] --- TODO: upgrade to a `UniformEmbedding` +-- TODO: upgrade to a `IsUniformEmbedding` theorem embedding_coeFn : Embedding (UniformFun.ofFun ∘ (⇑) : (α →ᵇ β) → α →ᵤ β) := ⟨inducing_coeFn, fun _ _ h => ext fun x => congr_fun h x⟩ @@ -524,7 +524,7 @@ theorem arzela_ascoli₂ (s : Set β) (hs : IsCompact s) (A : Set (α →ᵇ β) fun f hf => ?_ · haveI : CompactSpace s := isCompact_iff_compactSpace.1 hs refine arzela_ascoli₁ _ (continuous_iff_isClosed.1 (continuous_comp M) _ closed) ?_ - rw [uniformEmbedding_subtype_val.toUniformInducing.equicontinuous_iff] + rw [isUniformEmbedding_subtype_val.toUniformInducing.equicontinuous_iff] exact H.comp (A.restrictPreimage F) · let g := codRestrict s f fun x => in_s f x hf rw [show f = F g by ext; rfl] at hf ⊢ diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index eb72857ce4102..f68ff46c4011d 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -59,9 +59,12 @@ theorem uniformInducing_equivBoundedOfCompact : UniformInducing (equivBoundedOfC ⟨⟨Set.univ, { p | dist p.1 p.2 ≤ ε }⟩, ⟨isCompact_univ, ⟨ε, hε, fun _ h => h⟩⟩, fun ⟨f, g⟩ h => hs _ _ (ht ((dist_le hε.le).mpr fun x => h x (mem_univ x)))⟩⟩) -theorem uniformEmbedding_equivBoundedOfCompact : UniformEmbedding (equivBoundedOfCompact α β) := +theorem isUniformEmbedding_equivBoundedOfCompact : IsUniformEmbedding (equivBoundedOfCompact α β) := { uniformInducing_equivBoundedOfCompact α β with inj := (equivBoundedOfCompact α β).injective } +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_equivBoundedOfCompact := isUniformEmbedding_equivBoundedOfCompact + /-- When `α` is compact, the bounded continuous maps `α →ᵇ 𝕜` are additively equivalent to `C(α, 𝕜)`. -/ @@ -83,7 +86,7 @@ theorem addEquivBoundedOfCompact_apply [AddMonoid β] [LipschitzAdd β] : rfl instance metricSpace : MetricSpace C(α, β) := - (uniformEmbedding_equivBoundedOfCompact α β).comapMetricSpace _ + (isUniformEmbedding_equivBoundedOfCompact α β).comapMetricSpace _ /-- When `α` is compact, and `β` is a metric space, the bounded continuous maps `α →ᵇ β` are isometric to `C(α, β)`. diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index fd74c5bccfa7c..7dcfc9a8fff16 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -268,20 +268,26 @@ variable [Zero R] [UniformSpace R] protected instance instUniformSpace : UniformSpace C(X, R)₀ := .comap toContinuousMap inferInstance -lemma uniformEmbedding_toContinuousMap : - UniformEmbedding ((↑) : C(X, R)₀ → C(X, R)) where +lemma isUniformEmbedding_toContinuousMap : + IsUniformEmbedding ((↑) : C(X, R)₀ → C(X, R)) where comap_uniformity := rfl inj _ _ h := ext fun x ↦ congr($(h) x) +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toContinuousMap := isUniformEmbedding_toContinuousMap + instance [T1Space R] [CompleteSpace C(X, R)] : CompleteSpace C(X, R)₀ := - completeSpace_iff_isComplete_range uniformEmbedding_toContinuousMap.toUniformInducing + completeSpace_iff_isComplete_range isUniformEmbedding_toContinuousMap.toUniformInducing |>.mpr closedEmbedding_toContinuousMap.isClosed_range.isComplete -lemma uniformEmbedding_comp {Y : Type*} [UniformSpace Y] [Zero Y] (g : C(Y, R)₀) - (hg : UniformEmbedding g) : UniformEmbedding (g.comp · : C(X, Y)₀ → C(X, R)₀) := - uniformEmbedding_toContinuousMap.of_comp_iff.mp <| - ContinuousMap.uniformEmbedding_comp g.toContinuousMap hg |>.comp - uniformEmbedding_toContinuousMap +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)₀) := + isUniformEmbedding_toContinuousMap.of_comp_iff.mp <| + ContinuousMap.isUniformEmbedding_comp g.toContinuousMap hg |>.comp + isUniformEmbedding_toContinuousMap + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_comp := isUniformEmbedding_comp /-- The uniform equivalence `C(X, R)₀ ≃ᵤ C(Y, R)₀` induced by a homeomorphism of the domains sending `0 : X` to `0 : Y`. -/ @@ -291,12 +297,12 @@ def _root_.UniformEquiv.arrowCongrLeft₀ {Y : Type*} [TopologicalSpace Y] [Zero invFun g := g.comp ⟨f.toContinuousMap, hf⟩ left_inv g := ext fun _ ↦ congrArg g <| f.left_inv _ right_inv g := ext fun _ ↦ congrArg g <| f.right_inv _ - uniformContinuous_toFun := uniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| + uniformContinuous_toFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| ContinuousMap.uniformContinuous_comp_left f.symm.toContinuousMap |>.comp - uniformEmbedding_toContinuousMap.uniformContinuous - uniformContinuous_invFun := uniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| + isUniformEmbedding_toContinuousMap.uniformContinuous + uniformContinuous_invFun := isUniformEmbedding_toContinuousMap.uniformContinuous_iff.mpr <| ContinuousMap.uniformContinuous_comp_left f.toContinuousMap |>.comp - uniformEmbedding_toContinuousMap.uniformContinuous + isUniformEmbedding_toContinuousMap.uniformContinuous end UniformSpace @@ -340,7 +346,7 @@ section Norm variable {α : Type*} {𝕜 : Type*} {R : Type*} [TopologicalSpace α] [CompactSpace α] [Zero α] noncomputable instance [MetricSpace R] [Zero R]: MetricSpace C(α, R)₀ := - ContinuousMapZero.uniformEmbedding_toContinuousMap.comapMetricSpace _ + ContinuousMapZero.isUniformEmbedding_toContinuousMap.comapMetricSpace _ noncomputable instance [NormedAddCommGroup R] : Norm C(α, R)₀ where norm f := ‖(f : C(α, R))‖ diff --git a/Mathlib/Topology/EMetricSpace/Basic.lean b/Mathlib/Topology/EMetricSpace/Basic.lean index 20efe03e7d47a..deca954dd5665 100644 --- a/Mathlib/Topology/EMetricSpace/Basic.lean +++ b/Mathlib/Topology/EMetricSpace/Basic.lean @@ -66,20 +66,27 @@ theorem uniformInducing_iff [PseudoEMetricSpace β] {f : α → β} : simp only [subset_def, Prod.forall]; rfl /-- ε-δ characterization of uniform embeddings on pseudoemetric spaces -/ -nonrec theorem uniformEmbedding_iff [PseudoEMetricSpace β] {f : α → β} : - UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ +nonrec theorem isUniformEmbedding_iff [PseudoEMetricSpace β] {f : α → β} : + IsUniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < ε → edist a b < δ := - (uniformEmbedding_iff _).trans <| and_comm.trans <| Iff.rfl.and uniformInducing_iff + (isUniformEmbedding_iff _).trans <| and_comm.trans <| Iff.rfl.and uniformInducing_iff + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_iff := isUniformEmbedding_iff /-- If a map between pseudoemetric spaces is a uniform embedding then the edistance between `f x` and `f y` is controlled in terms of the distance between `x` and `y`. In fact, this lemma holds for a `UniformInducing` map. TODO: generalize? -/ -theorem controlled_of_uniformEmbedding [PseudoEMetricSpace β] {f : α → β} (h : UniformEmbedding f) : +theorem controlled_of_isUniformEmbedding [PseudoEMetricSpace β] {f : α → β} + (h : IsUniformEmbedding f) : (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, edist a b < δ → edist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < ε → edist a b < δ := - ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ + ⟨uniformContinuous_iff.1 h.uniformContinuous, (isUniformEmbedding_iff.1 h).2.2⟩ + +@[deprecated (since := "2024-10-01")] +alias controlled_of_uniformEmbedding := controlled_of_isUniformEmbedding /-- ε-δ characterization of Cauchy sequences on pseudoemetric spaces -/ protected theorem cauchy_iff {f : Filter α} : @@ -231,11 +238,14 @@ instance (priority := 100) EMetricSpace.instT0Space : T0Space γ where /-- A map between emetric spaces is a uniform embedding if and only if the edistance between `f x` and `f y` is controlled in terms of the distance between `x` and `y` and conversely. -/ -theorem EMetric.uniformEmbedding_iff' [EMetricSpace β] {f : γ → β} : - UniformEmbedding f ↔ +theorem EMetric.isUniformEmbedding_iff' [EMetricSpace β] {f : γ → β} : + IsUniformEmbedding f ↔ (∀ ε > 0, ∃ δ > 0, ∀ {a b : γ}, edist a b < δ → edist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : γ}, edist (f a) (f b) < ε → edist a b < δ := by - rw [uniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + rw [isUniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + +@[deprecated (since := "2024-10-01")] +alias EMetric.uniformEmbedding_iff' := EMetric.isUniformEmbedding_iff' /-- If a `PseudoEMetricSpace` is a T₀ space, then it is an `EMetricSpace`. -/ -- Porting note: made `reducible`; diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 12ff875ac5f87..2b6096cad74f3 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -25,7 +25,7 @@ to `EMetricSpace` at the end. -/ assert_not_exists Nat.instLocallyFiniteOrder -assert_not_exists UniformEmbedding +assert_not_exists IsUniformEmbedding assert_not_exists TendstoUniformlyOnFilter open Set Filter diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 4e93d2e862ea6..5ee25d7b929f2 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -39,13 +39,16 @@ theorem pairwise_one_le_dist : Pairwise fun m n : ℤ => 1 ≤ dist m n := by intro m n hne rw [dist_eq]; norm_cast; rwa [← zero_add (1 : ℤ), Int.add_one_le_iff, abs_pos, sub_ne_zero] -theorem uniformEmbedding_coe_real : UniformEmbedding ((↑) : ℤ → ℝ) := - uniformEmbedding_bot_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℤ → ℝ) := + isUniformEmbedding_bot_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[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 -instance : MetricSpace ℤ := Int.uniformEmbedding_coe_real.comapMetricSpace _ +instance : MetricSpace ℤ := Int.isUniformEmbedding_coe_real.comapMetricSpace _ theorem preimage_ball (x : ℤ) (r : ℝ) : (↑) ⁻¹' ball (x : ℝ) r = ball x r := rfl diff --git a/Mathlib/Topology/Instances/Nat.lean b/Mathlib/Topology/Instances/Nat.lean index 2e4857faf7c51..e90ce1cbe65b7 100644 --- a/Mathlib/Topology/Instances/Nat.lean +++ b/Mathlib/Topology/Instances/Nat.lean @@ -31,13 +31,16 @@ theorem dist_cast_real (x y : ℕ) : dist (x : ℝ) y = dist x y := rfl theorem pairwise_one_le_dist : Pairwise fun m n : ℕ => 1 ≤ dist m n := fun _ _ hne => Int.pairwise_one_le_dist <| mod_cast hne -theorem uniformEmbedding_coe_real : UniformEmbedding ((↑) : ℕ → ℝ) := - uniformEmbedding_bot_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist +theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℕ → ℝ) := + isUniformEmbedding_bot_of_pairwise_le_dist zero_lt_one pairwise_one_le_dist + +@[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 -instance : MetricSpace ℕ := Nat.uniformEmbedding_coe_real.comapMetricSpace _ +instance : MetricSpace ℕ := Nat.isUniformEmbedding_coe_real.comapMetricSpace _ theorem preimage_ball (x : ℕ) (r : ℝ) : (↑) ⁻¹' ball (x : ℝ) r = ball x r := rfl diff --git a/Mathlib/Topology/Instances/PNat.lean b/Mathlib/Topology/Instances/PNat.lean index 3a05c74ac7076..dd79778f35d9b 100644 --- a/Mathlib/Topology/Instances/PNat.lean +++ b/Mathlib/Topology/Instances/PNat.lean @@ -24,7 +24,10 @@ theorem dist_eq (x y : ℕ+) : dist x y = |(↑x : ℝ) - ↑y| := rfl @[simp, norm_cast] theorem dist_coe (x y : ℕ+) : dist (↑x : ℕ) (↑y : ℕ) = dist x y := rfl -theorem uniformEmbedding_coe : UniformEmbedding ((↑) : ℕ+ → ℕ) := uniformEmbedding_subtype_val +theorem isUniformEmbedding_coe : IsUniformEmbedding ((↑) : ℕ+ → ℕ) := isUniformEmbedding_subtype_val + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_coe := isUniformEmbedding_coe instance : DiscreteTopology ℕ+ := inferInstanceAs (DiscreteTopology { n : ℕ // 0 < n }) diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index 22cc2d6c6f682..a5585543b41fe 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -30,11 +30,14 @@ theorem dist_cast (x y : ℚ) : dist (x : ℝ) y = dist x y := theorem uniformContinuous_coe_real : UniformContinuous ((↑) : ℚ → ℝ) := uniformContinuous_comap -theorem uniformEmbedding_coe_real : UniformEmbedding ((↑) : ℚ → ℝ) := - uniformEmbedding_comap Rat.cast_injective +theorem isUniformEmbedding_coe_real : IsUniformEmbedding ((↑) : ℚ → ℝ) := + isUniformEmbedding_comap Rat.cast_injective + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_coe_real := isUniformEmbedding_coe_real theorem isDenseEmbedding_coe_real : IsDenseEmbedding ((↑) : ℚ → ℝ) := - uniformEmbedding_coe_real.isDenseEmbedding Rat.denseRange_cast + isUniformEmbedding_coe_real.isDenseEmbedding Rat.denseRange_cast @[deprecated (since := "2024-09-30")] alias denseEmbedding_coe_real := isDenseEmbedding_coe_real @@ -51,8 +54,11 @@ end Rat theorem Nat.dist_cast_rat (x y : ℕ) : dist (x : ℚ) y = dist x y := by rw [← Nat.dist_cast_real, ← Rat.dist_cast]; congr -theorem Nat.uniformEmbedding_coe_rat : UniformEmbedding ((↑) : ℕ → ℚ) := - uniformEmbedding_bot_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist +theorem Nat.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℕ → ℚ) := + isUniformEmbedding_bot_of_pairwise_le_dist zero_lt_one <| by simpa using Nat.pairwise_one_le_dist + +@[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 @@ -61,8 +67,11 @@ theorem Nat.closedEmbedding_coe_rat : ClosedEmbedding ((↑) : ℕ → ℚ) := theorem Int.dist_cast_rat (x y : ℤ) : dist (x : ℚ) y = dist x y := by rw [← Int.dist_cast_real, ← Rat.dist_cast]; congr -theorem Int.uniformEmbedding_coe_rat : UniformEmbedding ((↑) : ℤ → ℚ) := - uniformEmbedding_bot_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist +theorem Int.isUniformEmbedding_coe_rat : IsUniformEmbedding ((↑) : ℤ → ℚ) := + isUniformEmbedding_bot_of_pairwise_le_dist zero_lt_one <| by simpa using Int.pairwise_one_le_dist + +@[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 @@ -72,7 +81,7 @@ namespace Rat instance : NoncompactSpace ℚ := Int.closedEmbedding_coe_rat.noncompactSpace theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p.2 := - Rat.uniformEmbedding_coe_real.toUniformInducing.uniformContinuous_iff.2 <| by + Rat.isUniformEmbedding_coe_real.toUniformInducing.uniformContinuous_iff.2 <| by simp only [Function.comp_def, Rat.cast_add] exact Real.uniformContinuous_add.comp (Rat.uniformContinuous_coe_real.prod_map Rat.uniformContinuous_coe_real) @@ -97,7 +106,7 @@ instance : TopologicalRing ℚ := inferInstance nonrec theorem totallyBounded_Icc (a b : ℚ) : TotallyBounded (Icc a b) := by simpa only [preimage_cast_Icc] - using totallyBounded_preimage Rat.uniformEmbedding_coe_real.toUniformInducing + using totallyBounded_preimage Rat.isUniformEmbedding_coe_real.toUniformInducing (totallyBounded_Icc (a : ℝ) b) end Rat diff --git a/Mathlib/Topology/MetricSpace/Antilipschitz.lean b/Mathlib/Topology/MetricSpace/Antilipschitz.lean index 88e3996a46af3..ac87aa46237ea 100644 --- a/Mathlib/Topology/MetricSpace/Antilipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Antilipschitz.lean @@ -147,11 +147,12 @@ protected theorem uniformInducing (hf : AntilipschitzWith K f) (hfc : UniformCon UniformInducing f := ⟨le_antisymm hf.comap_uniformity_le hfc.le_comap⟩ -protected theorem uniformEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [PseudoEMetricSpace β] - {K : ℝ≥0} {f : α → β} (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : - UniformEmbedding f := +lemma isUniformEmbedding {α β : Type*} [EMetricSpace α] [PseudoEMetricSpace β] {K : ℝ≥0} {f : α → β} + (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : IsUniformEmbedding f := ⟨hf.uniformInducing hfc, hf.injective⟩ +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding + theorem isComplete_range [CompleteSpace α] (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : IsComplete (range f) := (hf.uniformInducing hfc).isComplete_range @@ -164,7 +165,7 @@ theorem isClosed_range {α β : Type*} [PseudoEMetricSpace α] [EMetricSpace β] theorem closedEmbedding {α : Type*} {β : Type*} [EMetricSpace α] [EMetricSpace β] {K : ℝ≥0} {f : α → β} [CompleteSpace α] (hf : AntilipschitzWith K f) (hfc : UniformContinuous f) : ClosedEmbedding f := - { (hf.uniformEmbedding hfc).embedding with isClosed_range := hf.isClosed_range hfc } + { (hf.isUniformEmbedding hfc).embedding with isClosed_range := hf.isClosed_range hfc } 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 f3b2468b6e50b..028c1ab133161 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -32,11 +32,14 @@ instance (priority := 100) _root_.MetricSpace.instT0Space : T0Space γ where /-- A map between metric spaces is a uniform embedding if and only if the distance between `f x` and `f y` is controlled in terms of the distance between `x` and `y` and conversely. -/ -theorem uniformEmbedding_iff' [MetricSpace β] {f : γ → β} : - UniformEmbedding f ↔ +theorem isUniformEmbedding_iff' [MetricSpace β] {f : γ → β} : + IsUniformEmbedding f ↔ (∀ ε > 0, ∃ δ > 0, ∀ {a b : γ}, dist a b < δ → dist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : γ}, dist (f a) (f b) < ε → dist a b < δ := by - rw [uniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + rw [isUniformEmbedding_iff_uniformInducing, uniformInducing_iff, uniformContinuous_iff] + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_iff' := isUniformEmbedding_iff' /-- If a `PseudoMetricSpace` is a T₀ space, then it is a `MetricSpace`. -/ abbrev _root_.MetricSpace.ofT0PseudoMetricSpace (α : Type*) [PseudoMetricSpace α] [T0Space α] : @@ -60,10 +63,13 @@ theorem closedEmbedding_of_pairwise_le_dist {α : Type*} [TopologicalSpace α] [ /-- 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 `β`. -/ -theorem uniformEmbedding_bot_of_pairwise_le_dist {β : Type*} {ε : ℝ} (hε : 0 < ε) {f : β → α} +theorem isUniformEmbedding_bot_of_pairwise_le_dist {β : Type*} {ε : ℝ} (hε : 0 < ε) {f : β → α} (hf : Pairwise fun x y => ε ≤ dist (f x) (f y)) : - @UniformEmbedding _ _ ⊥ (by infer_instance) f := - uniformEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + @IsUniformEmbedding _ _ ⊥ (by infer_instance) f := + isUniformEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_bot_of_pairwise_le_dist := isUniformEmbedding_bot_of_pairwise_le_dist end Metric @@ -94,10 +100,13 @@ abbrev MetricSpace.induced {γ β} (f : γ → β) (hf : Function.Injective f) ( /-- Pull back a metric space structure by a uniform embedding. This is a version of `MetricSpace.induced` useful in case if the domain already has a `UniformSpace` structure. -/ -abbrev UniformEmbedding.comapMetricSpace {α β} [UniformSpace α] [m : MetricSpace β] (f : α → β) - (h : UniformEmbedding f) : MetricSpace α := +abbrev IsUniformEmbedding.comapMetricSpace {α β} [UniformSpace α] [m : MetricSpace β] (f : α → β) + (h : IsUniformEmbedding f) : MetricSpace α := .replaceUniformity (.induced f h.inj m) h.comap_uniformity.symm +@[deprecated (since := "2024-10-03")] +alias UniformEmbedding.comapMetricSpace := IsUniformEmbedding.comapMetricSpace + /-- Pull back a metric space structure by an embedding. This is a version of `MetricSpace.induced` useful in case if the domain already has a `TopologicalSpace` structure. -/ abbrev Embedding.comapMetricSpace {α β} [TopologicalSpace α] [m : MetricSpace β] (f : α → β) diff --git a/Mathlib/Topology/MetricSpace/Closeds.lean b/Mathlib/Topology/MetricSpace/Closeds.lean index 67b9d4cbf3244..cf5140b7380c9 100644 --- a/Mathlib/Topology/MetricSpace/Closeds.lean +++ b/Mathlib/Topology/MetricSpace/Closeds.lean @@ -232,9 +232,12 @@ instance NonemptyCompacts.emetricSpace : EMetricSpace (NonemptyCompacts α) wher rwa [s.isCompact.isClosed.closure_eq, t.isCompact.isClosed.closure_eq] at this /-- `NonemptyCompacts.toCloseds` is a uniform embedding (as it is an isometry) -/ -theorem NonemptyCompacts.ToCloseds.uniformEmbedding : - UniformEmbedding (@NonemptyCompacts.toCloseds α _ _) := - Isometry.uniformEmbedding fun _ _ => rfl +theorem NonemptyCompacts.ToCloseds.isUniformEmbedding : + IsUniformEmbedding (@NonemptyCompacts.toCloseds α _ _) := + Isometry.isUniformEmbedding fun _ _ => rfl + +@[deprecated (since := "2024-10-01")] +alias NonemptyCompacts.ToCloseds.uniformEmbedding := NonemptyCompacts.ToCloseds.isUniformEmbedding /-- The range of `NonemptyCompacts.toCloseds` is closed in a complete space -/ theorem NonemptyCompacts.isClosed_in_closeds [CompleteSpace α] : @@ -278,14 +281,14 @@ theorem NonemptyCompacts.isClosed_in_closeds [CompleteSpace α] : from the same statement for closed subsets -/ instance NonemptyCompacts.completeSpace [CompleteSpace α] : CompleteSpace (NonemptyCompacts α) := (completeSpace_iff_isComplete_range - NonemptyCompacts.ToCloseds.uniformEmbedding.toUniformInducing).2 <| + NonemptyCompacts.ToCloseds.isUniformEmbedding.toUniformInducing).2 <| NonemptyCompacts.isClosed_in_closeds.isComplete /-- In a compact space, the type of nonempty compact subsets is compact. This follows from the same statement for closed subsets -/ instance NonemptyCompacts.compactSpace [CompactSpace α] : CompactSpace (NonemptyCompacts α) := ⟨by - rw [NonemptyCompacts.ToCloseds.uniformEmbedding.embedding.isCompact_iff, image_univ] + rw [NonemptyCompacts.ToCloseds.isUniformEmbedding.embedding.isCompact_iff, image_univ] exact NonemptyCompacts.isClosed_in_closeds.isCompact⟩ /-- In a second countable space, the type of nonempty compact subsets is second countable -/ diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index c68a245846375..b94867562a459 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -420,14 +420,16 @@ variable [EMetricSpace α] variable [FunLike F α β] /-- A dilation from a metric space is a uniform embedding -/ -protected theorem uniformEmbedding [PseudoEMetricSpace β] [DilationClass F α β] (f : F) : - UniformEmbedding f := - (antilipschitz f).uniformEmbedding (lipschitz f).uniformContinuous +lemma isUniformEmbedding [PseudoEMetricSpace β] [DilationClass F α β] (f : F) : + IsUniformEmbedding f := + (antilipschitz f).isUniformEmbedding (lipschitz f).uniformContinuous + +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding /-- A dilation from a metric space is an embedding -/ protected theorem embedding [PseudoEMetricSpace β] [DilationClass F α β] (f : F) : Embedding (f : α → β) := - (Dilation.uniformEmbedding f).embedding + (Dilation.isUniformEmbedding f).embedding /-- A dilation from a complete emetric space is a closed embedding -/ protected theorem closedEmbedding [CompleteSpace α] [EMetricSpace β] [DilationClass F α β] (f : F) : diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index fc48af3ac67e9..8b05b577ab5bf 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -162,12 +162,14 @@ protected theorem injective (h : Isometry f) : Injective f := h.antilipschitz.injective /-- An isometry from an emetric space is a uniform embedding -/ -protected theorem uniformEmbedding (hf : Isometry f) : UniformEmbedding f := - hf.antilipschitz.uniformEmbedding hf.lipschitz.uniformContinuous +lemma isUniformEmbedding (hf : Isometry f) : IsUniformEmbedding f := + hf.antilipschitz.isUniformEmbedding hf.lipschitz.uniformContinuous + +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding /-- An isometry from an emetric space is an embedding -/ protected theorem embedding (hf : Isometry f) : Embedding f := - hf.uniformEmbedding.embedding + hf.isUniformEmbedding.embedding /-- An isometry from a complete emetric space is a closed embedding -/ theorem closedEmbedding [CompleteSpace α] [EMetricSpace γ] {f : α → γ} (hf : Isometry f) : @@ -226,11 +228,14 @@ end Isometry -- namespace /-- A uniform embedding from a uniform space to a metric space is an isometry with respect to the induced metric space structure on the source space. -/ -theorem UniformEmbedding.to_isometry {α β} [UniformSpace α] [MetricSpace β] {f : α → β} - (h : UniformEmbedding f) : (letI := h.comapMetricSpace f; Isometry f) := +theorem IsUniformEmbedding.to_isometry {α β} [UniformSpace α] [MetricSpace β] {f : α → β} + (h : IsUniformEmbedding f) : (letI := h.comapMetricSpace f; Isometry f) := let _ := h.comapMetricSpace f Isometry.of_dist_eq fun _ _ => rfl +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.to_isometry := IsUniformEmbedding.to_isometry + /-- An embedding from a topological space to a metric space is an isometry with respect to the induced metric space structure on the source space. -/ theorem Embedding.to_isometry {α β} [TopologicalSpace α] [MetricSpace β] {f : α → β} diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean index 65722ebb2e6b9..346e0e97d3e23 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean @@ -68,17 +68,24 @@ nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod.map_apply, mem_setOf] -nonrec theorem uniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : - UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ +nonrec theorem isUniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : + IsUniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by - rw [uniformEmbedding_iff, and_comm, uniformInducing_iff] + rw [isUniformEmbedding_iff, and_comm, uniformInducing_iff] + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_iff := isUniformEmbedding_iff /-- If a map between pseudometric spaces is a uniform embedding then the distance between `f x` and `f y` is controlled in terms of the distance between `x` and `y`. -/ -theorem controlled_of_uniformEmbedding [PseudoMetricSpace β] {f : α → β} (h : UniformEmbedding f) : +theorem controlled_of_isUniformEmbedding [PseudoMetricSpace β] {f : α → β} + (h : IsUniformEmbedding f) : (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε) ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ + ⟨uniformContinuous_iff.1 h.uniformContinuous, (isUniformEmbedding_iff.1 h).2.2⟩ + +@[deprecated (since := "2024-10-01")] +alias controlled_of_uniformEmbedding := controlled_of_isUniformEmbedding theorem totallyBounded_iff {s : Set α} : TotallyBounded s ↔ ∀ ε > 0, ∃ t : Set α, t.Finite ∧ s ⊆ ⋃ y ∈ t, ball y ε := diff --git a/Mathlib/Topology/UniformSpace/Ascoli.lean b/Mathlib/Topology/UniformSpace/Ascoli.lean index 75e7105788ec2..0b702ae3a22f7 100644 --- a/Mathlib/Topology/UniformSpace/Ascoli.lean +++ b/Mathlib/Topology/UniformSpace/Ascoli.lean @@ -495,7 +495,8 @@ theorem ArzelaAscoli.isCompact_of_equicontinuous rw [isCompact_iff_compactSpace] at hS1 ⊢ exact (Equiv.toHomeomorphOfInducing _ h).symm.compactSpace rw [← inducing_subtype_val.of_comp_iff, ← EquicontinuousOn.inducing_uniformOnFun_iff_pi _ _ _] - · exact ContinuousMap.uniformEmbedding_toUniformOnFunIsCompact.inducing.comp inducing_subtype_val + · exact ContinuousMap.isUniformEmbedding_toUniformOnFunIsCompact.inducing.comp + inducing_subtype_val · exact eq_univ_iff_forall.mpr (fun x ↦ mem_sUnion_of_mem (mem_singleton x) isCompact_singleton) · exact fun _ ↦ id · exact fun K _ ↦ hS2.equicontinuousOn K diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 4d5a67fa923b1..d2dc3a7e09b3f 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -17,7 +17,7 @@ generalize to uniform spaces, e.g. * uniform continuity (in this file) * completeness (in `Cauchy.lean`) -* extension of uniform continuous functions to complete spaces (in `UniformEmbedding.lean`) +* extension of uniform continuous functions to complete spaces (in `IsUniformEmbedding.lean`) * totally bounded sets (in `Cauchy.lean`) * totally bounded complete sets are compact (in `Cauchy.lean`) diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index 7f50b07c0cc1d..c316d5316d094 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -171,11 +171,14 @@ instance compactConvergenceUniformSpace : UniformSpace C(α, β) := nhds_induced, tendsto_comap_iff, UniformOnFun.tendsto_iff_tendstoUniformlyOn] rfl -theorem uniformEmbedding_toUniformOnFunIsCompact : - UniformEmbedding (toUniformOnFunIsCompact : C(α, β) → α →ᵤ[{K | IsCompact K}] β) where +theorem isUniformEmbedding_toUniformOnFunIsCompact : + IsUniformEmbedding (toUniformOnFunIsCompact : C(α, β) → α →ᵤ[{K | IsCompact K}] β) where comap_uniformity := rfl inj := DFunLike.coe_injective +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_toUniformOnFunIsCompact := isUniformEmbedding_toUniformOnFunIsCompact + -- The following definitions and theorems -- used to be a part of the construction of the `UniformSpace C(α, β)` structure -- before it was migrated to `UniformOnFun` @@ -184,7 +187,7 @@ theorem _root_.Filter.HasBasis.compactConvergenceUniformity {ι : Type*} {pi : {s : ι → Set (β × β)} (h : (𝓤 β).HasBasis pi s) : HasBasis (𝓤 C(α, β)) (fun p : Set α × ι => IsCompact p.1 ∧ pi p.2) fun p => { fg : C(α, β) × C(α, β) | ∀ x ∈ p.1, (fg.1 x, fg.2 x) ∈ s p.2 } := by - rw [← uniformEmbedding_toUniformOnFunIsCompact.comap_uniformity] + rw [← isUniformEmbedding_toUniformOnFunIsCompact.comap_uniformity] exact .comap _ <| UniformOnFun.hasBasis_uniformity_of_basis _ _ {K | IsCompact K} ⟨∅, isCompact_empty⟩ (directedOn_of_sup_mem fun _ _ ↦ IsCompact.union) h @@ -260,27 +263,30 @@ variable {γ δ : Type*} [TopologicalSpace γ] [UniformSpace δ] theorem uniformContinuous_comp (g : C(β, δ)) (hg : UniformContinuous g) : UniformContinuous (ContinuousMap.comp g : C(α, β) → C(α, δ)) := - uniformEmbedding_toUniformOnFunIsCompact.uniformContinuous_iff.mpr <| + isUniformEmbedding_toUniformOnFunIsCompact.uniformContinuous_iff.mpr <| UniformOnFun.postcomp_uniformContinuous hg |>.comp - uniformEmbedding_toUniformOnFunIsCompact.uniformContinuous + isUniformEmbedding_toUniformOnFunIsCompact.uniformContinuous theorem uniformInducing_comp (g : C(β, δ)) (hg : UniformInducing g) : UniformInducing (ContinuousMap.comp g : C(α, β) → C(α, δ)) := - uniformEmbedding_toUniformOnFunIsCompact.toUniformInducing.of_comp_iff.mp <| + isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing.of_comp_iff.mp <| UniformOnFun.postcomp_uniformInducing hg |>.comp - uniformEmbedding_toUniformOnFunIsCompact.toUniformInducing + isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing + +theorem isUniformEmbedding_comp (g : C(β, δ)) (hg : IsUniformEmbedding g) : + IsUniformEmbedding (ContinuousMap.comp g : C(α, β) → C(α, δ)) := + isUniformEmbedding_toUniformOnFunIsCompact.of_comp_iff.mp <| + UniformOnFun.postcomp_isUniformEmbedding hg |>.comp + isUniformEmbedding_toUniformOnFunIsCompact -theorem uniformEmbedding_comp (g : C(β, δ)) (hg : UniformEmbedding g) : - UniformEmbedding (ContinuousMap.comp g : C(α, β) → C(α, δ)) := - uniformEmbedding_toUniformOnFunIsCompact.of_comp_iff.mp <| - UniformOnFun.postcomp_uniformEmbedding hg |>.comp - uniformEmbedding_toUniformOnFunIsCompact +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_comp := isUniformEmbedding_comp theorem uniformContinuous_comp_left (g : C(α, γ)) : UniformContinuous (fun f ↦ f.comp g : C(γ, β) → C(α, β)) := - uniformEmbedding_toUniformOnFunIsCompact.uniformContinuous_iff.mpr <| + isUniformEmbedding_toUniformOnFunIsCompact.uniformContinuous_iff.mpr <| UniformOnFun.precomp_uniformContinuous (fun _ hK ↦ hK.image g.continuous) |>.comp - uniformEmbedding_toUniformOnFunIsCompact.uniformContinuous + isUniformEmbedding_toUniformOnFunIsCompact.uniformContinuous /-- Any pair of a homeomorphism `X ≃ₜ Z` and an isomorphism `Y ≃ᵤ T` of uniform spaces gives rise to an isomorphism `C(X, Y) ≃ᵤ C(Z, T)`. -/ @@ -372,7 +378,8 @@ Sufficient conditions on `α` to satisfy this condition are (weak) local compact `ContinuousMap.instCompleteSpaceOfSequentialSpace`). -/ lemma completeSpace_of_restrictGenTopology (h : RestrictGenTopology {K : Set α | IsCompact K}) : CompleteSpace C(α, β) := by - rw [completeSpace_iff_isComplete_range uniformEmbedding_toUniformOnFunIsCompact.toUniformInducing, + rw [completeSpace_iff_isComplete_range + isUniformEmbedding_toUniformOnFunIsCompact.toUniformInducing, range_toUniformOnFunIsCompact, ← completeSpace_coe_iff_isComplete] exact (UniformOnFun.isClosed_setOf_continuous h).completeSpace_coe diff --git a/Mathlib/Topology/UniformSpace/CompareReals.lean b/Mathlib/Topology/UniformSpace/CompareReals.lean index 7a90e73b4afcc..7f1c135164444 100644 --- a/Mathlib/Topology/UniformSpace/CompareReals.lean +++ b/Mathlib/Topology/UniformSpace/CompareReals.lean @@ -72,7 +72,7 @@ def rationalCauSeqPkg : @AbstractCompletion ℚ <| (@AbsoluteValue.abs ℚ _).un (separation := by infer_instance) (uniformInducing := by rw [Rat.uniformSpace_eq] - exact Rat.uniformEmbedding_coe_real.toUniformInducing) + exact Rat.isUniformEmbedding_coe_real.toUniformInducing) (dense := Rat.isDenseEmbedding_coe_real.dense) namespace CompareReals diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 4bfc074965c87..2fcb34197b931 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -27,11 +27,14 @@ 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 UniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] - [T0Space β] {f : α → β} (hf : UniformEmbedding f) : +theorem IsUniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] + [T0Space β] {f : α → β} (hf : IsUniformEmbedding f) : ClosedEmbedding f := ⟨hf.embedding, hf.toUniformInducing.isComplete_range.isClosed⟩ +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.toClosedEmbedding := IsUniformEmbedding.toClosedEmbedding + namespace IsDenseInducing open Filter diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 0f386364453d4..2ca3f3f72a834 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -154,10 +154,13 @@ theorem uniformInducing_pureCauchy : UniformInducing (pureCauchy : α → Cauchy _ = 𝓤 α := by simp [this] ⟩ -theorem uniformEmbedding_pureCauchy : UniformEmbedding (pureCauchy : α → CauchyFilter α) := +theorem isUniformEmbedding_pureCauchy : IsUniformEmbedding (pureCauchy : α → CauchyFilter α) := { uniformInducing_pureCauchy with inj := fun _a₁ _a₂ h => pure_injective <| Subtype.ext_iff_val.1 h } +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_pureCauchy := isUniformEmbedding_pureCauchy + theorem denseRange_pureCauchy : DenseRange (pureCauchy : α → CauchyFilter α) := fun f => by have h_ex : ∀ s ∈ 𝓤 (CauchyFilter α), ∃ y : α, (f, pureCauchy y) ∈ s := fun s hs => let ⟨t'', ht''₁, (ht''₂ : gen t'' ⊆ s)⟩ := (mem_lift'_sets monotone_gen).mp hs @@ -184,7 +187,7 @@ theorem isDenseInducing_pureCauchy : IsDenseInducing (pureCauchy : α → Cauchy uniformInducing_pureCauchy.isDenseInducing denseRange_pureCauchy theorem isDenseEmbedding_pureCauchy : IsDenseEmbedding (pureCauchy : α → CauchyFilter α) := - uniformEmbedding_pureCauchy.isDenseEmbedding denseRange_pureCauchy + isUniformEmbedding_pureCauchy.isDenseEmbedding denseRange_pureCauchy @[deprecated (since := "2024-09-30")] alias denseEmbedding_pureCauchy := isDenseEmbedding_pureCauchy @@ -359,12 +362,15 @@ theorem uniformContinuous_coe : UniformContinuous ((↑) : α → Completion α) theorem continuous_coe : Continuous ((↑) : α → Completion α) := cPkg.continuous_coe -theorem uniformEmbedding_coe [T0Space α] : UniformEmbedding ((↑) : α → Completion α) := +theorem isUniformEmbedding_coe [T0Space α] : IsUniformEmbedding ((↑) : α → Completion α) := { comap_uniformity := comap_coe_eq_uniformity α inj := separated_pureCauchy_injective } +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_coe := isUniformEmbedding_coe + theorem coe_injective [T0Space α] : Function.Injective ((↑) : α → Completion α) := - UniformEmbedding.inj (uniformEmbedding_coe _) + IsUniformEmbedding.inj (isUniformEmbedding_coe _) variable {α} diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index d4be20cb620ad..a52b28256431d 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -203,14 +203,16 @@ protected theorem uniformInducing (h : α ≃ᵤ β) : UniformInducing h := theorem comap_eq (h : α ≃ᵤ β) : UniformSpace.comap h ‹_› = ‹_› := h.uniformInducing.comap_uniformSpace -protected theorem uniformEmbedding (h : α ≃ᵤ β) : UniformEmbedding h := - ⟨h.uniformInducing, h.injective⟩ +theorem isUniformEmbedding (h : α ≃ᵤ β) : IsUniformEmbedding h := ⟨h.uniformInducing, h.injective⟩ + +@[deprecated (since := "2024-10-01")] alias uniformEmbedding := isUniformEmbedding theorem completeSpace_iff (h : α ≃ᵤ β) : CompleteSpace α ↔ CompleteSpace β := - completeSpace_congr h.uniformEmbedding + completeSpace_congr h.isUniformEmbedding /-- Uniform equiv given a uniform embedding. -/ -noncomputable def ofUniformEmbedding (f : α → β) (hf : UniformEmbedding f) : α ≃ᵤ Set.range f where +noncomputable def ofIsUniformEmbedding (f : α → β) (hf : IsUniformEmbedding f) : + α ≃ᵤ Set.range f where uniformContinuous_toFun := hf.toUniformInducing.uniformContinuous.subtype_mk _ uniformContinuous_invFun := by rw [hf.toUniformInducing.uniformContinuous_iff, Equiv.invFun_as_coe, @@ -218,6 +220,8 @@ noncomputable def ofUniformEmbedding (f : α → β) (hf : UniformEmbedding f) : exact uniformContinuous_subtype_val toEquiv := Equiv.ofInjective f hf.inj +@[deprecated (since := "2024-10-03")] alias ofUniformEmbedding := ofIsUniformEmbedding + /-- If two sets are equal, then they are uniformly equivalent. -/ def setCongr {s t : Set α} (h : s = t) : s ≃ᵤ t where uniformContinuous_toFun := uniformContinuous_subtype_val.subtype_mk _ diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 4eae4a2c7b4eb..851e5d68e405e 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -377,11 +377,15 @@ a uniform embedding for the uniform structures of uniform convergence. More precisely, if `f : γ → β` is a uniform embedding, then `(f ∘ ·) : (α →ᵤ γ) → (α →ᵤ β)` is a uniform embedding. -/ -protected theorem postcomp_uniformEmbedding [UniformSpace γ] {f : γ → β} (hf : UniformEmbedding f) : - UniformEmbedding (ofFun ∘ (f ∘ ·) ∘ toFun : (α →ᵤ γ) → α →ᵤ β) where +protected theorem postcomp_isUniformEmbedding [UniformSpace γ] {f : γ → β} + (hf : IsUniformEmbedding f) : + IsUniformEmbedding (ofFun ∘ (f ∘ ·) ∘ toFun : (α →ᵤ γ) → α →ᵤ β) where toUniformInducing := UniformFun.postcomp_uniformInducing hf.toUniformInducing inj _ _ H := funext fun _ ↦ hf.inj (congrFun H _) +@[deprecated (since := "2024-10-01")] +alias postcomp_uniformEmbedding := UniformFun.postcomp_isUniformEmbedding + -- Porting note: had to add a type annotation at `((f ∘ ·) : ((α → γ) → (α → β)))` /-- If `u` is a uniform structures on `β` and `f : γ → β`, then `𝒰(α, γ, comap f u) = comap (fun g ↦ f ∘ g) 𝒰(α, γ, u₁)`. -/ @@ -877,11 +881,14 @@ uniform structures of `𝔖`-convergence. More precisely, if `f : γ → β` is a uniform embedding, then `(fun g ↦ f ∘ g) : (α →ᵤ[𝔖] γ) → (α →ᵤ[𝔖] β)` is a uniform embedding. -/ -protected theorem postcomp_uniformEmbedding [UniformSpace γ] {f : γ → β} (hf : UniformEmbedding f) : - UniformEmbedding (ofFun 𝔖 ∘ (f ∘ ·) ∘ toFun 𝔖) where +protected theorem postcomp_isUniformEmbedding [UniformSpace γ] {f : γ → β} + (hf : IsUniformEmbedding f) : IsUniformEmbedding (ofFun 𝔖 ∘ (f ∘ ·) ∘ toFun 𝔖) where toUniformInducing := UniformOnFun.postcomp_uniformInducing hf.toUniformInducing inj _ _ H := funext fun _ ↦ hf.inj (congrFun H _) +@[deprecated (since := "2024-10-01")] +alias postcomp_uniformEmbedding := UniformOnFun.postcomp_isUniformEmbedding + /-- Turn a uniform isomorphism `γ ≃ᵤ β` into a uniform isomorphism `(α →ᵤ[𝔖] γ) ≃ᵤ (α →ᵤ[𝔖] β)` by post-composing. -/ protected def congrRight [UniformSpace γ] (e : γ ≃ᵤ β) : (α →ᵤ[𝔖] γ) ≃ᵤ (α →ᵤ[𝔖] β) := diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 048ea735a5df4..adad8bcab4c26 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -27,7 +27,7 @@ variable {α : Type u} {β : Type v} {γ : Type w} [UniformSpace α] [UniformSpa /-- A map `f : α → β` between uniform spaces is called *uniform inducing* if the uniformity filter on `α` is the pullback of the uniformity filter on `β` under `Prod.map f f`. If `α` is a separated -space, then this implies that `f` is injective, hence it is a `UniformEmbedding`. -/ +space, then this implies that `f` is injective, hence it is a `IsUniformEmbedding`. -/ @[mk_iff] structure UniformInducing (f : α → β) : Prop where /-- The uniformity filter on the domain is the pullback of the uniformity filter on the codomain @@ -130,70 +130,109 @@ protected theorem UniformInducing.injective [T0Space α] {f : α → β} (h : Un /-- A map `f : α → β` between uniform spaces is a *uniform embedding* if it is uniform inducing and injective. If `α` is a separated space, then the latter assumption follows from the former. -/ @[mk_iff] -structure UniformEmbedding (f : α → β) extends UniformInducing f : Prop where +structure IsUniformEmbedding (f : α → β) extends UniformInducing f : Prop where /-- A uniform embedding is injective. -/ inj : Function.Injective f -theorem uniformEmbedding_iff' {f : α → β} : - UniformEmbedding f ↔ Injective f ∧ UniformContinuous f ∧ comap (Prod.map f f) (𝓤 β) ≤ 𝓤 α := by - rw [uniformEmbedding_iff, and_comm, uniformInducing_iff'] +@[deprecated (since := "2024-10-03")] alias UniformEmbedding := IsUniformEmbedding -theorem Filter.HasBasis.uniformEmbedding_iff' {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} +theorem isUniformEmbedding_iff' {f : α → β} : + IsUniformEmbedding f ↔ + Injective f ∧ UniformContinuous f ∧ comap (Prod.map f f) (𝓤 β) ≤ 𝓤 α := by + rw [isUniformEmbedding_iff, and_comm, uniformInducing_iff'] + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_iff' := isUniformEmbedding_iff' + +theorem Filter.HasBasis.isUniformEmbedding_iff' {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} (h : (𝓤 α).HasBasis p s) (h' : (𝓤 β).HasBasis p' s') {f : α → β} : - UniformEmbedding f ↔ Injective f ∧ + IsUniformEmbedding f ↔ Injective f ∧ (∀ i, p' i → ∃ j, p j ∧ ∀ x y, (x, y) ∈ s j → (f x, f y) ∈ s' i) ∧ (∀ j, p j → ∃ i, p' i ∧ ∀ x y, (f x, f y) ∈ s' i → (x, y) ∈ s j) := by - rw [uniformEmbedding_iff, and_comm, h.uniformInducing_iff h'] + rw [isUniformEmbedding_iff, and_comm, h.uniformInducing_iff h'] + +@[deprecated (since := "2024-10-01")] +alias Filter.HasBasis.uniformEmbedding_iff' := Filter.HasBasis.isUniformEmbedding_iff' -theorem Filter.HasBasis.uniformEmbedding_iff {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} +theorem Filter.HasBasis.isUniformEmbedding_iff {ι ι'} {p : ι → Prop} {p' : ι' → Prop} {s s'} (h : (𝓤 α).HasBasis p s) (h' : (𝓤 β).HasBasis p' s') {f : α → β} : - UniformEmbedding f ↔ Injective f ∧ UniformContinuous f ∧ + IsUniformEmbedding f ↔ Injective f ∧ UniformContinuous f ∧ (∀ j, p j → ∃ i, p' i ∧ ∀ x y, (f x, f y) ∈ s' i → (x, y) ∈ s j) := by - simp only [h.uniformEmbedding_iff' h', h.uniformContinuous_iff h'] + simp only [h.isUniformEmbedding_iff' h', h.uniformContinuous_iff h'] -theorem uniformEmbedding_subtype_val {p : α → Prop} : - UniformEmbedding (Subtype.val : Subtype p → α) := +@[deprecated (since := "2024-10-01")] +alias Filter.HasBasis.uniformEmbedding_iff := Filter.HasBasis.isUniformEmbedding_iff + +theorem isUniformEmbedding_subtype_val {p : α → Prop} : + IsUniformEmbedding (Subtype.val : Subtype p → α) := { comap_uniformity := rfl inj := Subtype.val_injective } -theorem uniformEmbedding_set_inclusion {s t : Set α} (hst : s ⊆ t) : - UniformEmbedding (inclusion hst) where +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_subtype_val := isUniformEmbedding_subtype_val + +theorem isUniformEmbedding_set_inclusion {s t : Set α} (hst : s ⊆ t) : + IsUniformEmbedding (inclusion hst) where comap_uniformity := by rw [uniformity_subtype, uniformity_subtype, comap_comap]; rfl inj := inclusion_injective hst -theorem UniformEmbedding.comp {g : β → γ} (hg : UniformEmbedding g) {f : α → β} - (hf : UniformEmbedding f) : UniformEmbedding (g ∘ f) := +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_set_inclusion := isUniformEmbedding_set_inclusion + +theorem IsUniformEmbedding.comp {g : β → γ} (hg : IsUniformEmbedding g) {f : α → β} + (hf : IsUniformEmbedding f) : IsUniformEmbedding (g ∘ f) := { hg.toUniformInducing.comp hf.toUniformInducing with inj := hg.inj.comp hf.inj } -theorem UniformEmbedding.of_comp_iff {g : β → γ} (hg : UniformEmbedding g) {f : α → β} : - UniformEmbedding (g ∘ f) ↔ UniformEmbedding f := by - simp_rw [uniformEmbedding_iff, hg.toUniformInducing.of_comp_iff, hg.inj.of_comp_iff f] +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.comp := IsUniformEmbedding.comp -theorem Equiv.uniformEmbedding {α β : Type*} [UniformSpace α] [UniformSpace β] (f : α ≃ β) - (h₁ : UniformContinuous f) (h₂ : UniformContinuous f.symm) : UniformEmbedding f := - uniformEmbedding_iff'.2 ⟨f.injective, h₁, by rwa [← Equiv.prodCongr_apply, ← map_equiv_symm]⟩ +theorem IsUniformEmbedding.of_comp_iff {g : β → γ} (hg : IsUniformEmbedding g) {f : α → β} : + IsUniformEmbedding (g ∘ f) ↔ IsUniformEmbedding f := by + simp_rw [isUniformEmbedding_iff, hg.toUniformInducing.of_comp_iff, hg.inj.of_comp_iff f] -theorem uniformEmbedding_inl : UniformEmbedding (Sum.inl : α → α ⊕ β) := - uniformEmbedding_iff'.2 ⟨Sum.inl_injective, uniformContinuous_inl, fun s hs => +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.of_comp_iff := IsUniformEmbedding.of_comp_iff + +theorem Equiv.isUniformEmbedding {α β : Type*} [UniformSpace α] [UniformSpace β] (f : α ≃ β) + (h₁ : UniformContinuous f) (h₂ : UniformContinuous f.symm) : IsUniformEmbedding f := + isUniformEmbedding_iff'.2 ⟨f.injective, h₁, by rwa [← Equiv.prodCongr_apply, ← map_equiv_symm]⟩ + +@[deprecated (since := "2024-10-01")] +alias Equiv.uniformEmbedding := Equiv.isUniformEmbedding + +theorem isUniformEmbedding_inl : IsUniformEmbedding (Sum.inl : α → α ⊕ β) := + isUniformEmbedding_iff'.2 ⟨Sum.inl_injective, uniformContinuous_inl, fun s hs => ⟨Prod.map Sum.inl Sum.inl '' s ∪ range (Prod.map Sum.inr Sum.inr), union_mem_sup (image_mem_map hs) range_mem_map, fun x h => by simpa [Prod.map_apply'] using h⟩⟩ -theorem uniformEmbedding_inr : UniformEmbedding (Sum.inr : β → α ⊕ β) := - uniformEmbedding_iff'.2 ⟨Sum.inr_injective, uniformContinuous_inr, fun s hs => +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_inl := isUniformEmbedding_inl + +theorem isUniformEmbedding_inr : IsUniformEmbedding (Sum.inr : β → α ⊕ β) := + isUniformEmbedding_iff'.2 ⟨Sum.inr_injective, uniformContinuous_inr, fun s hs => ⟨range (Prod.map Sum.inl Sum.inl) ∪ Prod.map Sum.inr Sum.inr '' s, union_mem_sup range_mem_map (image_mem_map hs), fun x h => by simpa [Prod.map_apply'] using h⟩⟩ +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_inr := isUniformEmbedding_inr + /-- If the domain of a `UniformInducing` map `f` is a T₀ space, then `f` is injective, -hence it is a `UniformEmbedding`. -/ -protected theorem UniformInducing.uniformEmbedding [T0Space α] {f : α → β} - (hf : UniformInducing f) : UniformEmbedding f := +hence it is a `IsUniformEmbedding`. -/ +protected theorem UniformInducing.isUniformEmbedding [T0Space α] {f : α → β} + (hf : UniformInducing f) : IsUniformEmbedding f := ⟨hf, hf.inducing.injective⟩ -theorem uniformEmbedding_iff_uniformInducing [T0Space α] {f : α → β} : - UniformEmbedding f ↔ UniformInducing f := - ⟨UniformEmbedding.toUniformInducing, UniformInducing.uniformEmbedding⟩ +@[deprecated (since := "2024-10-01")] +alias UniformInducing.uniformEmbedding := UniformInducing.isUniformEmbedding + +theorem isUniformEmbedding_iff_uniformInducing [T0Space α] {f : α → β} : + IsUniformEmbedding f ↔ UniformInducing f := + ⟨IsUniformEmbedding.toUniformInducing, UniformInducing.isUniformEmbedding⟩ + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_iff_uniformInducing := isUniformEmbedding_iff_uniformInducing /-- If a map `f : α → β` sends any two distinct points to point that are **not** related by a fixed `s ∈ 𝓤 β`, then `f` is uniform inducing with respect to the discrete uniformity on `α`: @@ -210,28 +249,37 @@ theorem comap_uniformity_of_spaced_out {α} {f : α → β} {s : Set (β × β)} /-- If a map `f : α → β` sends any two distinct points to point that are **not** related by a fixed `s ∈ 𝓤 β`, then `f` is a uniform embedding with respect to the discrete uniformity on `α`. -/ -theorem uniformEmbedding_of_spaced_out {α} {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) - (hf : Pairwise fun x y => (f x, f y) ∉ s) : @UniformEmbedding α β ⊥ ‹_› f := by +theorem isUniformEmbedding_of_spaced_out {α} {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) + (hf : Pairwise fun x y => (f x, f y) ∉ s) : @IsUniformEmbedding α β ⊥ ‹_› f := by let _ : UniformSpace α := ⊥; have := discreteTopology_bot α - exact UniformInducing.uniformEmbedding ⟨comap_uniformity_of_spaced_out hs hf⟩ + exact UniformInducing.isUniformEmbedding ⟨comap_uniformity_of_spaced_out hs hf⟩ + +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_of_spaced_out := isUniformEmbedding_of_spaced_out -protected theorem UniformEmbedding.embedding {f : α → β} (h : UniformEmbedding f) : Embedding f := +protected lemma IsUniformEmbedding.embedding {f : α → β} (h : IsUniformEmbedding f) : Embedding f := { toInducing := h.toUniformInducing.inducing inj := h.inj } -theorem UniformEmbedding.isDenseEmbedding {f : α → β} (h : UniformEmbedding f) (hd : DenseRange f) : - IsDenseEmbedding f := +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.embedding := IsUniformEmbedding.embedding + +theorem IsUniformEmbedding.isDenseEmbedding {f : α → β} (h : IsUniformEmbedding f) + (hd : DenseRange f) : IsDenseEmbedding f := { h.embedding with dense := hd } +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.isDenseEmbedding := IsUniformEmbedding.isDenseEmbedding + @[deprecated (since := "2024-09-30")] -alias UniformEmbedding.denseEmbedding := UniformEmbedding.isDenseEmbedding +alias IsUniformEmbedding.denseEmbedding := IsUniformEmbedding.isDenseEmbedding theorem closedEmbedding_of_spaced_out {α} [TopologicalSpace α] [DiscreteTopology α] [T0Space β] {f : α → β} {s : Set (β × β)} (hs : s ∈ 𝓤 β) (hf : Pairwise fun x y => (f x, f y) ∉ s) : ClosedEmbedding f := by rcases @DiscreteTopology.eq_bot α _ _ with rfl; let _ : UniformSpace α := ⊥ exact - { (uniformEmbedding_of_spaced_out hs hf).embedding with + { (isUniformEmbedding_of_spaced_out hs hf).embedding with isClosed_range := isClosed_range_of_spaced_out hs hf } theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α → β} (b : β) @@ -248,18 +296,24 @@ theorem closure_image_mem_nhds_of_uniformInducing {s : Set (α × α)} {e : α rcases he₂.dense.mem_nhds (inter_mem hV (ho.mem_nhds hy)) with ⟨x, hxV, hxU⟩ exact ⟨e x, hxV, mem_image_of_mem e hxU⟩ -theorem uniformEmbedding_subtypeEmb (p : α → Prop) {e : α → β} (ue : UniformEmbedding e) - (de : IsDenseEmbedding e) : UniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := +theorem isUniformEmbedding_subtypeEmb (p : α → Prop) {e : α → β} (ue : IsUniformEmbedding e) + (de : IsDenseEmbedding e) : IsUniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := { comap_uniformity := by simp [comap_comap, Function.comp_def, IsDenseEmbedding.subtypeEmb, uniformity_subtype, ue.comap_uniformity.symm] inj := (de.subtype p).inj } -theorem UniformEmbedding.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [UniformSpace β'] - {e₁ : α → α'} {e₂ : β → β'} (h₁ : UniformEmbedding e₁) (h₂ : UniformEmbedding e₂) : - UniformEmbedding fun p : α × β => (e₁ p.1, e₂ p.2) := +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_subtypeEmb := isUniformEmbedding_subtypeEmb + +theorem IsUniformEmbedding.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [UniformSpace β'] + {e₁ : α → α'} {e₂ : β → β'} (h₁ : IsUniformEmbedding e₁) (h₂ : IsUniformEmbedding e₂) : + IsUniformEmbedding fun p : α × β => (e₁ p.1, e₂ p.2) := { h₁.toUniformInducing.prod h₂.toUniformInducing with inj := h₁.inj.prodMap h₂.inj } +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.prod := IsUniformEmbedding.prod + /-- A set is complete iff its image under a uniform inducing map is complete. -/ theorem isComplete_image_iff {m : α → β} {s : Set α} (hm : UniformInducing m) : IsComplete (m '' s) ↔ IsComplete s := by @@ -273,15 +327,18 @@ theorem isComplete_image_iff {m : α → β} {s : Set α} (hm : UniformInducing theorem UniformInducing.isComplete_iff {f : α → β} {s : Set α} (hf : UniformInducing f) : IsComplete (f '' s) ↔ IsComplete s := isComplete_image_iff hf -/-- If `f : X → Y` is an `UniformEmbedding`, the image `f '' s` of a set `s` is complete +/-- If `f : X → Y` is an `IsUniformEmbedding`, the image `f '' s` of a set `s` is complete if and only if `s` is complete. -/ -theorem UniformEmbedding.isComplete_iff {f : α → β} {s : Set α} (hf : UniformEmbedding f) : +theorem IsUniformEmbedding.isComplete_iff {f : α → β} {s : Set α} (hf : IsUniformEmbedding f) : IsComplete (f '' s) ↔ IsComplete s := hf.toUniformInducing.isComplete_iff +@[deprecated (since := "2024-10-01")] +alias UniformEmbedding.isComplete_iff := IsUniformEmbedding.isComplete_iff + /-- Sets of a subtype are complete iff their image under the coercion is complete. -/ theorem Subtype.isComplete_iff {p : α → Prop} {s : Set { x // p x }} : IsComplete s ↔ IsComplete ((↑) '' s : Set α) := - uniformEmbedding_subtype_val.isComplete_iff.symm + isUniformEmbedding_subtype_val.isComplete_iff.symm alias ⟨isComplete_of_complete_image, _⟩ := isComplete_image_iff @@ -312,12 +369,12 @@ instance SeparationQuotient.instCompleteSpace [CompleteSpace α] : /-- See also `UniformInducing.completeSpace_congr` for a version that works for non-injective maps. -/ -theorem completeSpace_congr {e : α ≃ β} (he : UniformEmbedding e) : +theorem completeSpace_congr {e : α ≃ β} (he : IsUniformEmbedding e) : CompleteSpace α ↔ CompleteSpace β := he.completeSpace_congr e.surjective theorem completeSpace_coe_iff_isComplete {s : Set α} : CompleteSpace s ↔ IsComplete s := by - rw [completeSpace_iff_isComplete_range uniformEmbedding_subtype_val.toUniformInducing, + rw [completeSpace_iff_isComplete_range isUniformEmbedding_subtype_val.toUniformInducing, Subtype.range_coe] alias ⟨_, IsComplete.completeSpace_coe⟩ := completeSpace_coe_iff_isComplete @@ -394,28 +451,34 @@ theorem totallyBounded_preimage {f : α → β} {s : Set β} (hf : UniformInduci instance CompleteSpace.sum [CompleteSpace α] [CompleteSpace β] : CompleteSpace (α ⊕ β) := by rw [completeSpace_iff_isComplete_univ, ← range_inl_union_range_inr] - exact uniformEmbedding_inl.toUniformInducing.isComplete_range.union - uniformEmbedding_inr.toUniformInducing.isComplete_range + exact isUniformEmbedding_inl.toUniformInducing.isComplete_range.union + isUniformEmbedding_inr.toUniformInducing.isComplete_range end -theorem uniformEmbedding_comap {α : Type*} {β : Type*} {f : α → β} [u : UniformSpace β] - (hf : Function.Injective f) : @UniformEmbedding α β (UniformSpace.comap f u) u f := - @UniformEmbedding.mk _ _ (UniformSpace.comap f u) _ _ +theorem isUniformEmbedding_comap {α : Type*} {β : Type*} {f : α → β} [u : UniformSpace β] + (hf : Function.Injective f) : @IsUniformEmbedding α β (UniformSpace.comap f u) u f := + @IsUniformEmbedding.mk _ _ (UniformSpace.comap f u) _ _ (@UniformInducing.mk _ _ (UniformSpace.comap f u) _ _ rfl) hf +@[deprecated (since := "2024-10-01")] +alias uniformEmbedding_comap := isUniformEmbedding_comap + /-- Pull back a uniform space structure by an embedding, adjusting the new uniform structure to make sure that its topology is defeq to the original one. -/ def Embedding.comapUniformSpace {α β} [TopologicalSpace α] [u : UniformSpace β] (f : α → β) (h : Embedding f) : UniformSpace α := (u.comap f).replaceTopology h.induced -theorem Embedding.to_uniformEmbedding {α β} [TopologicalSpace α] [u : UniformSpace β] (f : α → β) - (h : Embedding f) : @UniformEmbedding α β (h.comapUniformSpace f) u f := +theorem Embedding.to_isUniformEmbedding {α β} [TopologicalSpace α] [u : UniformSpace β] (f : α → β) + (h : Embedding f) : @IsUniformEmbedding α β (h.comapUniformSpace f) u f := let _ := h.comapUniformSpace f { comap_uniformity := rfl inj := h.inj } +@[deprecated (since := "2024-10-01")] +alias Embedding.to_uniformEmbedding := Embedding.to_isUniformEmbedding + section UniformExtension variable {α : Type*} {β : Type*} {γ : Type*} [UniformSpace α] [UniformSpace β] [UniformSpace γ] @@ -434,13 +497,13 @@ theorem uniformly_extend_exists [CompleteSpace γ] (a : α) : ∃ c, Tendsto f ( CompleteSpace.complete this theorem uniform_extend_subtype [CompleteSpace γ] {p : α → Prop} {e : α → β} {f : α → γ} {b : β} - {s : Set α} (hf : UniformContinuous fun x : Subtype p => f x.val) (he : UniformEmbedding e) + {s : Set α} (hf : UniformContinuous fun x : Subtype p => f x.val) (he : IsUniformEmbedding e) (hd : ∀ x : β, x ∈ closure (range e)) (hb : closure (e '' s) ∈ 𝓝 b) (hs : IsClosed s) (hp : ∀ x ∈ s, p x) : ∃ c, Tendsto f (comap e (𝓝 b)) (𝓝 c) := by have de : IsDenseEmbedding e := he.isDenseEmbedding hd have de' : IsDenseEmbedding (IsDenseEmbedding.subtypeEmb p e) := de.subtype p - have ue' : UniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := - uniformEmbedding_subtypeEmb _ he de + have ue' : IsUniformEmbedding (IsDenseEmbedding.subtypeEmb p e) := + isUniformEmbedding_subtypeEmb _ he de have : b ∈ closure (e '' { x | p x }) := (closure_mono <| monotone_image <| hp) (mem_of_mem_nhds hb) let ⟨c, hc⟩ := uniformly_extend_exists ue'.toUniformInducing de'.dense hf ⟨b, this⟩ diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 61a8f1beba401..de0b42c05c8e5 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -1315,7 +1315,7 @@ Filter.HasBasis.prod_nhds' Filter.HasBasis.sup' Filter.HasBasis.to_hasBasis' Filter.HasBasis.to_image_id' -Filter.HasBasis.uniformEmbedding_iff' +Filter.HasBasis.isUniformEmbedding_iff' Filter.iInf_neBot_iff_of_directed' Filter.iInf_sets_eq_finite' Filter.isScalarTower' @@ -4727,7 +4727,7 @@ uniformContinuous_mul_left' uniformContinuous_mul_right' uniformContinuous_nnnorm' uniformContinuous_norm' -uniformEmbedding_iff' +isUniformEmbedding_iff' UniformGroup.mk' uniformInducing_iff' UniformInducing.mk' From 1c37f1e9241238bc538b1e4c3921d117f650894b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 4 Oct 2024 03:33:32 +0000 Subject: [PATCH 165/174] chore: bump toolchain to v4.13.0-rc2 (#17377) This merges the already reviewed `bump/v4.13.0` branch. Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Johan Commelin --- .../IfNormalization/WithoutAesop.lean | 10 +- Cache/IO.lean | 2 +- Cache/Requests.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 1 - Mathlib/Algebra/BigOperators/Group/List.lean | 13 +-- Mathlib/Algebra/CharZero/Defs.lean | 1 - .../ContinuedFractions/ConvergentsEquiv.lean | 3 - Mathlib/Algebra/Field/Defs.lean | 4 +- Mathlib/Algebra/Free.lean | 1 - Mathlib/Algebra/Group/Hom/Basic.lean | 2 +- Mathlib/Algebra/Group/Hom/Defs.lean | 5 +- Mathlib/Algebra/Group/Subgroup/Basic.lean | 12 +-- Mathlib/Algebra/Group/Subgroup/Finite.lean | 8 +- .../Algebra/Group/Submonoid/Operations.lean | 2 +- .../Group/Subsemigroup/Operations.lean | 2 +- Mathlib/Algebra/Group/ZeroOne.lean | 13 +-- .../GroupWithZero/NonZeroDivisors.lean | 2 +- .../Homology/HomologicalBicomplex.lean | 2 +- Mathlib/Algebra/Homology/TotalComplex.lean | 4 +- .../Algebra/Homology/TotalComplexShift.lean | 4 +- Mathlib/Algebra/Module/Submodule/Lattice.lean | 2 +- Mathlib/Algebra/NeZero.lean | 26 +----- .../Order/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Order/Floor.lean | 22 ++++- Mathlib/Algebra/Order/Group/Abs.lean | 4 +- Mathlib/Algebra/Order/Group/Cone.lean | 4 +- Mathlib/Algebra/Order/Group/Defs.lean | 10 +- Mathlib/Algebra/Order/Hom/Basic.lean | 40 +++++--- Mathlib/Algebra/Order/Monoid/Prod.lean | 2 +- Mathlib/Algebra/Order/Ring/Cone.lean | 2 +- Mathlib/Algebra/Order/ZeroLEOne.lean | 1 - Mathlib/Algebra/Polynomial/BigOperators.lean | 2 +- .../Polynomial/Degree/CardPowDegree.lean | 2 +- Mathlib/Algebra/Ring/CentroidHom.lean | 4 +- Mathlib/Algebra/Ring/Parity.lean | 6 +- Mathlib/Algebra/Ring/Subring/Basic.lean | 2 +- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 4 +- Mathlib/Algebra/Star/Free.lean | 2 +- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 9 +- .../Morphisms/QuasiCompact.lean | 10 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 8 +- .../ProjectiveSpectrum/Scheme.lean | 9 +- Mathlib/AlgebraicGeometry/Pullbacks.lean | 5 +- Mathlib/Analysis/Analytic/Composition.lean | 4 +- .../Instances.lean | 12 +-- .../NonUnital.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 4 +- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 10 +- .../Analysis/InnerProductSpace/OfNorm.lean | 2 +- .../NormedSpace/Multilinear/Basic.lean | 2 +- .../Analysis/SpecialFunctions/Integrals.lean | 5 +- .../SpecialFunctions/Trigonometric/Basic.lean | 2 +- .../CategoryTheory/ConnectedComponents.lean | 1 - Mathlib/CategoryTheory/Extensive.lean | 4 +- .../Functor/KanExtension/Adjunction.lean | 2 +- .../CategoryTheory/GradedObject/Monoidal.lean | 2 +- .../LiftingProperties/Basic.lean | 2 +- .../CategoryTheory/Limits/Shapes/Images.lean | 2 - .../Limits/Shapes/Pullback/CommSq.lean | 16 ++-- .../MorphismProperty/Basic.lean | 16 ++-- .../Sites/Coherent/SheafComparison.lean | 20 +++- .../CategoryTheory/Sites/Grothendieck.lean | 8 +- .../CategoryTheory/Triangulated/Functor.lean | 2 +- .../Combinatorics/Additive/AP/Three/Defs.lean | 4 +- Mathlib/Combinatorics/Configuration.lean | 2 +- .../Combinatorics/Enumerative/DyckWord.lean | 5 +- .../SetFamily/KruskalKatona.lean | 12 ++- .../SimpleGraph/Hamiltonian.lean | 6 +- Mathlib/Combinatorics/SimpleGraph/Path.lean | 2 +- .../SimpleGraph/Regularity/Equitabilise.lean | 8 +- .../SimpleGraph/Triangle/Basic.lean | 6 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 3 +- Mathlib/Combinatorics/Young/YoungDiagram.lean | 2 +- Mathlib/Computability/Ackermann.lean | 6 +- Mathlib/Computability/Halting.lean | 4 +- Mathlib/Computability/Language.lean | 2 +- Mathlib/Computability/PartrecCode.lean | 8 +- Mathlib/Computability/TuringMachine.lean | 2 +- Mathlib/Control/Applicative.lean | 3 +- Mathlib/Control/Basic.lean | 11 +-- Mathlib/Control/Functor.lean | 4 +- Mathlib/Control/Traversable/Basic.lean | 1 + Mathlib/Control/Traversable/Equiv.lean | 2 +- Mathlib/Control/Traversable/Instances.lean | 4 +- Mathlib/Data/Array/ExtractLemmas.lean | 8 +- Mathlib/Data/DFinsupp/Order.lean | 9 +- Mathlib/Data/Fin/Basic.lean | 31 +------ Mathlib/Data/Fin/Tuple/Basic.lean | 2 +- Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 4 +- Mathlib/Data/Finset/Functor.lean | 9 +- Mathlib/Data/Finset/NatDivisors.lean | 6 +- Mathlib/Data/Fintype/Card.lean | 2 - Mathlib/Data/Fintype/Prod.lean | 2 - Mathlib/Data/Int/Defs.lean | 8 +- Mathlib/Data/List/Basic.lean | 92 +++++-------------- Mathlib/Data/List/Chain.lean | 24 ++++- Mathlib/Data/List/Cycle.lean | 4 +- Mathlib/Data/List/Dedup.lean | 2 +- Mathlib/Data/List/Defs.lean | 1 + Mathlib/Data/List/Forall2.lean | 8 +- Mathlib/Data/List/GetD.lean | 7 +- Mathlib/Data/List/Indexes.lean | 23 ++--- Mathlib/Data/List/Infix.lean | 15 ++- Mathlib/Data/List/InsertNth.lean | 6 +- Mathlib/Data/List/Intervals.lean | 4 +- Mathlib/Data/List/Lattice.lean | 2 +- Mathlib/Data/List/Lemmas.lean | 8 -- Mathlib/Data/List/MinMax.lean | 22 +++-- Mathlib/Data/List/Nodup.lean | 35 ++++--- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Perm.lean | 25 +---- Mathlib/Data/List/Permutation.lean | 6 +- Mathlib/Data/List/Rotate.lean | 10 +- Mathlib/Data/List/Sort.lean | 36 ++------ Mathlib/Data/List/Sublists.lean | 12 +-- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 29 ++++-- Mathlib/Data/Multiset/Functor.lean | 4 +- Mathlib/Data/Nat/BitIndices.lean | 4 +- Mathlib/Data/Nat/Bits.lean | 15 ++- Mathlib/Data/Nat/Bitwise.lean | 23 +---- Mathlib/Data/Nat/Cast/NeZero.lean | 1 - Mathlib/Data/Nat/Defs.lean | 18 ++-- Mathlib/Data/Nat/Digits.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 8 +- Mathlib/Data/Nat/Factorization/Defs.lean | 2 +- Mathlib/Data/Nat/Log.lean | 8 +- Mathlib/Data/Nat/WithBot.lean | 6 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Option/Basic.lean | 35 ------- Mathlib/Data/Option/Defs.lean | 4 - Mathlib/Data/Ordmap/Ordset.lean | 4 +- Mathlib/Data/PFunctor/Univariate/M.lean | 4 + Mathlib/Data/PNat/Defs.lean | 1 - Mathlib/Data/Prod/Basic.lean | 4 +- Mathlib/Data/Prod/Lex.lean | 4 +- Mathlib/Data/Rat/Lemmas.lean | 12 +-- Mathlib/Data/Seq/WSeq.lean | 7 +- Mathlib/Data/Set/MemPartition.lean | 1 - Mathlib/Data/Setoid/Basic.lean | 2 +- Mathlib/Data/Vector/Defs.lean | 2 +- Mathlib/Data/ZMod/Defs.lean | 2 +- Mathlib/Dynamics/PeriodicPts.lean | 4 +- .../IsAlgClosed/AlgebraicClosure.lean | 2 +- Mathlib/FieldTheory/KummerExtension.lean | 5 + Mathlib/FieldTheory/PurelyInseparable.lean | 2 +- Mathlib/FieldTheory/Separable.lean | 2 +- Mathlib/FieldTheory/SeparableDegree.lean | 2 +- .../Geometry/RingedSpace/OpenImmersion.lean | 44 ++++----- .../RingedSpace/PresheafedSpace/Gluing.lean | 31 +++---- Mathlib/GroupTheory/CommutingProbability.lean | 1 + Mathlib/GroupTheory/FreeGroup/Basic.lean | 5 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 6 ++ Mathlib/GroupTheory/OrderOfElement.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 2 +- Mathlib/GroupTheory/Perm/Sign.lean | 2 +- .../LinearAlgebra/Alternating/DomCoprod.lean | 2 +- .../CliffordAlgebra/Grading.lean | 6 +- Mathlib/LinearAlgebra/Dual.lean | 2 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 4 +- .../LinearAlgebra/Matrix/Charpoly/Coeff.lean | 2 +- .../Matrix/Determinant/Basic.lean | 4 +- Mathlib/LinearAlgebra/Semisimple.lean | 2 +- .../TensorProduct/Graded/Internal.lean | 2 +- Mathlib/Logic/Basic.lean | 2 +- Mathlib/Logic/Encodable/Basic.lean | 6 +- Mathlib/Logic/Equiv/Array.lean | 2 +- Mathlib/Logic/Nonempty.lean | 7 -- .../Constructions/BorelSpace/Order.lean | 12 +-- Mathlib/MeasureTheory/Measure/Content.lean | 2 +- Mathlib/ModelTheory/Algebra/Ring/Basic.lean | 2 +- Mathlib/ModelTheory/Basic.lean | 4 +- Mathlib/ModelTheory/Encoding.lean | 4 +- Mathlib/ModelTheory/Syntax.lean | 4 +- Mathlib/NumberTheory/ADEInequality.lean | 2 +- Mathlib/NumberTheory/Bertrand.lean | 6 +- .../NumberTheory/Cyclotomic/Discriminant.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 24 ++++- .../LegendreSymbol/GaussEisensteinLemmas.lean | 6 +- .../LegendreSymbol/JacobiSymbol.lean | 10 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/PythagoreanTriples.lean | 2 +- Mathlib/NumberTheory/SmoothNumbers.lean | 2 +- Mathlib/NumberTheory/SumFourSquares.lean | 9 +- Mathlib/Order/Directed.lean | 2 +- Mathlib/Order/Filter/Bases.lean | 4 +- Mathlib/Order/Hom/Basic.lean | 3 +- Mathlib/Order/Hom/Bounded.lean | 6 +- Mathlib/Order/Interval/Finset/Nat.lean | 8 +- Mathlib/Order/RelIso/Basic.lean | 2 +- .../Order/SuccPred/LinearLocallyFinite.lean | 9 +- Mathlib/Order/SymmDiff.lean | 2 +- .../ProbabilityMassFunction/Basic.lean | 3 +- .../ProbabilityMassFunction/Monad.lean | 6 +- .../GroupCohomology/LowDegree.lean | 2 +- Mathlib/RingTheory/AdjoinRoot.lean | 4 +- .../IsIntegralClosure/Basic.lean | 6 +- Mathlib/RingTheory/LocalRing/Module.lean | 26 +++--- .../RingTheory/Localization/FractionRing.lean | 2 +- .../Symmetric/NewtonIdentities.lean | 6 +- .../NonUnitalSubsemiring/Basic.lean | 2 +- .../Polynomial/Cyclotomic/Eval.lean | 9 +- Mathlib/RingTheory/RingHomProperties.lean | 6 +- Mathlib/RingTheory/SimpleRing/Basic.lean | 2 +- Mathlib/RingTheory/Smooth/StandardSmooth.lean | 4 +- Mathlib/SetTheory/Cardinal/Finite.lean | 8 +- Mathlib/Tactic/CC/Addition.lean | 8 +- Mathlib/Tactic/IntervalCases.lean | 1 + Mathlib/Tactic/NormNum/DivMod.lean | 6 +- Mathlib/Tactic/Ring/Basic.lean | 2 +- Mathlib/Tactic/WLOG.lean | 3 +- Mathlib/Testing/SlimCheck/Functions.lean | 2 +- .../Topology/Algebra/Valued/NormedValued.lean | 2 +- .../Category/LightProfinite/Extend.lean | 4 +- .../Topology/Category/Profinite/Extend.lean | 4 +- Mathlib/Topology/ContinuousMap/Basic.lean | 4 +- Mathlib/Topology/MetricSpace/Dilation.lean | 2 +- Mathlib/Topology/Metrizable/Uniformity.lean | 4 +- Mathlib/Topology/Order/LawsonTopology.lean | 2 +- Mathlib/Topology/UniformSpace/Cauchy.lean | 2 +- Mathlib/Util/CountHeartbeats.lean | 4 +- lake-manifest.json | 6 +- lakefile.lean | 3 +- lean-toolchain | 2 +- scripts/noshake.json | 3 +- test/ValuedCSP.lean | 2 + test/aesop_cat.lean | 2 + test/matrix.lean | 4 +- test/says.lean | 3 +- 232 files changed, 806 insertions(+), 758 deletions(-) diff --git a/Archive/Examples/IfNormalization/WithoutAesop.lean b/Archive/Examples/IfNormalization/WithoutAesop.lean index 1ac14973a5106..b3e94e0a5653e 100644 --- a/Archive/Examples/IfNormalization/WithoutAesop.lean +++ b/Archive/Examples/IfNormalization/WithoutAesop.lean @@ -92,8 +92,8 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) : · simp_all · have := ht₃ v have := he₃ v - simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_true', - AList.lookup_insert_eq_none, ne_eq, AList.lookup_insert] + simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_eq_eq_not, + Bool.not_true, AList.lookup_insert_eq_none, ne_eq, AList.lookup_insert] obtain ⟨⟨⟨tn, tc⟩, tr⟩, td⟩ := ht₂ split <;> rename_i h' · subst h' @@ -103,9 +103,9 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) : have := he₃ w by_cases h : w = v · subst h; simp_all - · simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_true', - AList.lookup_insert_eq_none, ne_eq, not_false_eq_true, AList.lookup_insert_ne, - implies_true] + · simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_eq_eq_not, + Bool.not_true, AList.lookup_insert_eq_none, ne_eq, not_false_eq_true, + AList.lookup_insert_ne, implies_true] obtain ⟨⟨⟨en, ec⟩, er⟩, ed⟩ := he₂ split at b <;> rename_i h' · subst h'; simp_all diff --git a/Cache/IO.lean b/Cache/IO.lean index 6cbd295c31031..43e3f70652623 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -338,7 +338,7 @@ def packCache (hashMap : HashMap) (overwrite verbose unpackedOnly : Bool) /-- Gets the set of all cached files -/ def getLocalCacheSet : IO <| Lean.RBTree String compare := do let paths ← getFilesWithExtension CACHEDIR "ltar" - return .fromList (paths.data.map (·.withoutParent CACHEDIR |>.toString)) _ + return .fromList (paths.toList.map (·.withoutParent CACHEDIR |>.toString)) _ def isPathFromMathlib (path : FilePath) : Bool := match path.components with diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 5c4c6038aecaa..71b74cb4b051d 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -184,7 +184,7 @@ def UPLOAD_URL : String := /-- Formats the config file for `curl`, containing the list of files to be uploaded -/ def mkPutConfigContent (fileNames : Array String) (token : String) : IO String := do let token := if useFROCache then "" else s!"?{token}" -- the FRO cache doesn't pass the token here - let l ← fileNames.data.mapM fun fileName : String => do + let l ← fileNames.toList.mapM fun fileName : String => do pure s!"-T {(IO.CACHEDIR / fileName).toString}\nurl = {mkFileURL UPLOAD_URL fileName}{token}" return "\n".intercalate l diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 9d3b4b23f9e67..ed65a54d0f15f 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -248,7 +248,6 @@ instance instAddTorsor : AddTorsor (G × G') (P × P') where zero_vadd _ := Prod.ext (zero_vadd _ _) (zero_vadd _ _) add_vadd _ _ _ := Prod.ext (add_vadd _ _ _) (add_vadd _ _ _) vsub p₁ p₂ := (p₁.1 -ᵥ p₂.1, p₁.2 -ᵥ p₂.2) - nonempty := Prod.instNonempty vsub_vadd' _ _ := Prod.ext (vsub_vadd _ _) (vsub_vadd _ _) vadd_vsub' _ _ := Prod.ext (vadd_vsub _ _) (vadd_vsub _ _) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 59be51355e078..6a1a5fb8af68e 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -125,7 +125,7 @@ theorem prod_replicate (n : ℕ) (a : M) : (replicate n a).prod = a ^ n := by @[to_additive sum_eq_card_nsmul] theorem prod_eq_pow_card (l : List M) (m : M) (h : ∀ x ∈ l, x = m) : l.prod = m ^ l.length := by - rw [← prod_replicate, ← List.eq_replicate.mpr ⟨rfl, h⟩] + rw [← prod_replicate, ← List.eq_replicate_iff.mpr ⟨rfl, h⟩] @[to_additive] theorem prod_hom_rel (l : List ι) {r : M → N → Prop} {f : ι → M} {g : ι → N} (h₁ : r 1 1) @@ -639,20 +639,15 @@ lemma ranges_join (l : List ℕ) : l.ranges.join = range l.sum := by simp [range lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by simp [mem_mem_ranges_iff_lt_natSum] -lemma countP_join (p : α → Bool) : ∀ L : List (List α), countP p L.join = (L.map (countP p)).sum - | [] => rfl - | a :: l => by rw [join, countP_append, map_cons, sum_cons, countP_join _ l] - -lemma count_join [BEq α] (L : List (List α)) (a : α) : L.join.count a = (L.map (count a)).sum := - countP_join _ _ - @[simp] theorem length_bind (l : List α) (f : α → List β) : length (List.bind l f) = sum (map (length ∘ f) l) := by rw [List.bind, length_join, map_map, Nat.sum_eq_listSum] lemma countP_bind (p : β → Bool) (l : List α) (f : α → List β) : - countP p (l.bind f) = sum (map (countP p ∘ f) l) := by rw [List.bind, countP_join, map_map] + countP p (l.bind f) = sum (map (countP p ∘ f) l) := by + rw [List.bind, countP_join, map_map] + simp lemma count_bind [BEq β] (l : List α) (f : α → List β) (x : β) : count x (l.bind f) = sum (map (count x ∘ f) l) := countP_bind _ _ _ diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index 8f2cf74d4cef2..c82356e6fa6e7 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Int.Cast.Defs -import Mathlib.Algebra.NeZero import Mathlib.Logic.Function.Basic /-! diff --git a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean index 2adf582909701..99e76fa64a5f5 100644 --- a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -166,9 +166,6 @@ theorem succ_succ_nth_conv'Aux_eq_succ_nth_conv'Aux_squashSeq : gp_head.a / (gp_head.b + convs'Aux s.tail (m + 2)) = convs'Aux (squashSeq s (m + 1)) (m + 2) by simpa only [convs'Aux, s_head_eq] - have : convs'Aux s.tail (m + 2) = convs'Aux (squashSeq s.tail m) (m + 1) := by - refine IH gp_succ_n ?_ - simpa [Stream'.Seq.get?_tail] using s_succ_nth_eq have : (squashSeq s (m + 1)).head = some gp_head := (squashSeq_nth_of_lt m.succ_pos).trans s_head_eq simp_all [convs'Aux, squashSeq_succ_n_tail_eq_squashSeq_tail_n] diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index 834f25d5fc450..46e889528e3e6 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -45,8 +45,8 @@ field, division ring, skew field, skew-field, skewfield assert_not_imported Mathlib.Tactic.Common --- `NeZero` should not be needed in the basic algebraic hierarchy. -assert_not_exists NeZero +-- `NeZero` theory should not be needed in the basic algebraic hierarchy +assert_not_imported Mathlib.Algebra.NeZero assert_not_exists MonoidHom diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index 1b0f927525cca..134a4e81ed9ff 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -6,7 +6,6 @@ Authors: Kenny Lau import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Control.Applicative import Mathlib.Control.Traversable.Basic -import Mathlib.Data.List.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Tactic.AdaptationNote diff --git a/Mathlib/Algebra/Group/Hom/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean index 148098efacb87..094fd70d45cc5 100644 --- a/Mathlib/Algebra/Group/Hom/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -14,7 +14,7 @@ import Mathlib.Algebra.Group.Hom.Defs -- `NeZero` cannot be additivised, hence its theory should be developed outside of the -- `Algebra.Group` folder. -assert_not_exists NeZero +assert_not_imported Mathlib.Algebra.NeZero variable {α β M N P : Type*} diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index febcba6ed781c..2279253985c32 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -143,8 +143,9 @@ homomorphisms. You should also extend this typeclass when you extend `AddMonoidHom`. -/ -class AddMonoidHomClass (F M N : Type*) [AddZeroClass M] [AddZeroClass N] [FunLike F M N] - extends AddHomClass F M N, ZeroHomClass F M N : Prop +class AddMonoidHomClass (F : Type*) (M N : outParam Type*) + [AddZeroClass M] [AddZeroClass N] [FunLike F M N] + extends AddHomClass F M N, ZeroHomClass F M N : Prop -- Instances and lemmas are defined below through `@[to_additive]`. end add_zero diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 2df0f53da5529..556df038ddca9 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -94,27 +94,27 @@ variable {A : Type*} [AddGroup A] section SubgroupClass /-- `InvMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ -class InvMemClass (S G : Type*) [Inv G] [SetLike S G] : Prop where +class InvMemClass (S : Type*) (G : outParam Type*) [Inv G] [SetLike S G] : Prop where /-- `s` is closed under inverses -/ inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s export InvMemClass (inv_mem) /-- `NegMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ -class NegMemClass (S G : Type*) [Neg G] [SetLike S G] : Prop where +class NegMemClass (S : Type*) (G : outParam Type*) [Neg G] [SetLike S G] : Prop where /-- `s` is closed under negation -/ neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s export NegMemClass (neg_mem) /-- `SubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ -class SubgroupClass (S G : Type*) [DivInvMonoid G] [SetLike S G] extends SubmonoidClass S G, - InvMemClass S G : Prop +class SubgroupClass (S : Type*) (G : outParam Type*) [DivInvMonoid G] [SetLike S G] + extends SubmonoidClass S G, InvMemClass S G : Prop /-- `AddSubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are additive subgroups of `G`. -/ -class AddSubgroupClass (S G : Type*) [SubNegMonoid G] [SetLike S G] extends AddSubmonoidClass S G, - NegMemClass S G : Prop +class AddSubgroupClass (S : Type*) (G : outParam Type*) [SubNegMonoid G] [SetLike S G] + extends AddSubmonoidClass S G, NegMemClass S G : Prop attribute [to_additive] InvMemClass SubgroupClass diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index b4dc8b267698d..6e5342a357f4e 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -165,9 +165,11 @@ theorem pi_mem_of_mulSingle_mem_aux [DecidableEq η] (I : Finset η) {H : Subgro x ∈ H := by induction I using Finset.induction_on generalizing x with | empty => - convert one_mem H - ext i - exact h1 i (Finset.not_mem_empty i) + have : x = 1 := by + ext i + exact h1 i (Finset.not_mem_empty i) + rw [this] + exact one_mem H | insert hnmem ih => rename_i i I have : x = Function.update x i 1 * Pi.mulSingle i (x i) := by diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 1bf4b67f90258..f84160423a5b5 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -591,7 +591,7 @@ theorem closure_closure_coe_preimage {s : Set M} : closure (((↑) : closure s Subtype.recOn x fun x hx _ => by refine closure_induction' (p := fun y hy ↦ (⟨y, hy⟩ : closure s) ∈ closure (((↑) : closure s → M) ⁻¹' s)) - (fun g hg => subset_closure hg) ?_ (fun g₁ g₂ hg₁ hg₂ => ?_) hx + _ (fun g hg => subset_closure hg) ?_ (fun g₁ g₂ hg₁ hg₂ => ?_) hx · exact Submonoid.one_mem _ · exact Submonoid.mul_mem _ diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index 05de353785319..85f523619d7ca 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -501,7 +501,7 @@ theorem closure_closure_coe_preimage {s : Set M} : eq_top_iff.2 fun x => Subtype.recOn x fun _ hx' _ => closure_induction' (p := fun y hy ↦ (⟨y, hy⟩ : closure s) ∈ closure (((↑) : closure s → M) ⁻¹' s)) - (fun _ hg => subset_closure hg) (fun _ _ _ _ => Subsemigroup.mul_mem _) hx' + _ (fun _ hg => subset_closure hg) (fun _ _ _ _ => Subsemigroup.mul_mem _) hx' /-- Given `Subsemigroup`s `s`, `t` of semigroups `M`, `N` respectively, `s × t` as a subsemigroup of `M × N`. -/ diff --git a/Mathlib/Algebra/Group/ZeroOne.lean b/Mathlib/Algebra/Group/ZeroOne.lean index 3e72968095621..8822f97d1d17e 100644 --- a/Mathlib/Algebra/Group/ZeroOne.lean +++ b/Mathlib/Algebra/Group/ZeroOne.lean @@ -6,17 +6,10 @@ Authors: Gabriel Ebner, Mario Carneiro import Mathlib.Tactic.ToAdditive /-! -## Classes for `Zero` and `One` --/ - -class Zero.{u} (α : Type u) where - zero : α +## Typeclass `One` -instance (priority := 300) Zero.toOfNat0 {α} [Zero α] : OfNat α (nat_lit 0) where - ofNat := ‹Zero α›.1 - -instance (priority := 200) Zero.ofOfNat0 {α} [OfNat α (nat_lit 0)] : Zero α where - zero := 0 +`Zero` has already been defined in Lean. +-/ universe u diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index 3c09328f414ca..2c4e1832cb5a3 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -304,7 +304,7 @@ theorem mk_mem_nonZeroDivisors_associates : Associates.mk a ∈ (Associates M₀ /-- The non-zero divisors of associates of a monoid with zero `M₀` are isomorphic to the associates of the non-zero divisors of `M₀` under the map `⟨⟦a⟧, _⟩ ↦ ⟦⟨a, _⟩⟧`. -/ def associatesNonZeroDivisorsEquiv : (Associates M₀)⁰ ≃* Associates M₀⁰ where - toEquiv := .subtypeQuotientEquivQuotientSubtype (s₂ := Associated.setoid _) + toEquiv := .subtypeQuotientEquivQuotientSubtype _ (s₂ := Associated.setoid _) (· ∈ nonZeroDivisors _) (by simp [mem_nonZeroDivisors_iff, Quotient.forall, Associates.mk_mul_mk]) (by simp [Associated.setoid]) diff --git a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean index b02aad6a648b8..dbd85f9d8e710 100644 --- a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean @@ -205,7 +205,7 @@ def XXIsoOfEq {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : @[simp] lemma XXIsoOfEq_rfl (i₁ : I₁) (i₂ : I₂) : - K.XXIsoOfEq (rfl : i₁ = i₁) (rfl : i₂ = i₂) = Iso.refl _ := rfl + K.XXIsoOfEq _ _ _ (rfl : i₁ = i₁) (rfl : i₂ = i₂) = Iso.refl _ := rfl end HomologicalComplex₂ diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 73fb3959181c0..645d747de064f 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -260,7 +260,7 @@ noncomputable def ιTotal (i₁ : I₁) (i₂ : I₂) (i₁₂ : I₁₂) @[reassoc (attr := simp)] lemma XXIsoOfEq_hom_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : x₂ = y₂) (i₁₂ : I₁₂) (h : ComplexShape.π c₁ c₂ c₁₂ (y₁, y₂) = i₁₂) : - (K.XXIsoOfEq h₁ h₂).hom ≫ K.ιTotal c₁₂ y₁ y₂ i₁₂ h = + (K.XXIsoOfEq _ _ _ h₁ h₂).hom ≫ K.ιTotal c₁₂ y₁ y₂ i₁₂ h = K.ιTotal c₁₂ x₁ x₂ i₁₂ (by rw [h₁, h₂, h]) := by subst h₁ h₂ simp @@ -268,7 +268,7 @@ lemma XXIsoOfEq_hom_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : @[reassoc (attr := simp)] lemma XXIsoOfEq_inv_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : x₂ = y₂) (i₁₂ : I₁₂) (h : ComplexShape.π c₁ c₂ c₁₂ (x₁, x₂) = i₁₂) : - (K.XXIsoOfEq h₁ h₂).inv ≫ K.ιTotal c₁₂ x₁ x₂ i₁₂ h = + (K.XXIsoOfEq _ _ _ h₁ h₂).inv ≫ K.ιTotal c₁₂ x₁ x₂ i₁₂ h = K.ιTotal c₁₂ y₁ y₂ i₁₂ (by rw [← h, h₁, h₂]) := by subst h₁ h₂ simp diff --git a/Mathlib/Algebra/Homology/TotalComplexShift.lean b/Mathlib/Algebra/Homology/TotalComplexShift.lean index 8d32945310df2..345677aa57934 100644 --- a/Mathlib/Algebra/Homology/TotalComplexShift.lean +++ b/Mathlib/Algebra/Homology/TotalComplexShift.lean @@ -129,7 +129,7 @@ noncomputable def totalShift₁XIso (n n' : ℤ) (h : n + x = n') : (((shiftFunctor₁ C x).obj K).total (up ℤ)).X n ≅ (K.total (up ℤ)).X n' where hom := totalDesc _ (fun p q hpq => K.ιTotal (up ℤ) (p + x) q n' (by dsimp at hpq ⊢; omega)) inv := totalDesc _ (fun p q hpq => - (K.XXIsoOfEq (Int.sub_add_cancel p x) rfl).inv ≫ + (K.XXIsoOfEq _ _ _ (Int.sub_add_cancel p x) rfl).inv ≫ ((shiftFunctor₁ C x).obj K).ιTotal (up ℤ) (p - x) q n (by dsimp at hpq ⊢; omega)) hom_inv_id := by @@ -235,7 +235,7 @@ noncomputable def totalShift₂XIso (n n' : ℤ) (h : n + y = n') : hom := totalDesc _ (fun p q hpq => (p * y).negOnePow • K.ιTotal (up ℤ) p (q + y) n' (by dsimp at hpq ⊢; omega)) inv := totalDesc _ (fun p q hpq => (p * y).negOnePow • - (K.XXIsoOfEq rfl (Int.sub_add_cancel q y)).inv ≫ + (K.XXIsoOfEq _ _ _ rfl (Int.sub_add_cancel q y)).inv ≫ ((shiftFunctor₂ C y).obj K).ιTotal (up ℤ) p (q - y) n (by dsimp at hpq ⊢; omega)) hom_inv_id := by ext p q h diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index ce617ef5b2a6b..3e15b4721bc18 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -291,7 +291,7 @@ theorem toAddSubmonoid_sSup (s : Set (Submodule R M)) : { toAddSubmonoid := sSup (toAddSubmonoid '' s) smul_mem' := fun t {m} h ↦ by simp_rw [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, sSup_eq_iSup'] at h ⊢ - refine AddSubmonoid.iSup_induction' + refine AddSubmonoid.iSup_induction' _ (C := fun x _ ↦ t • x ∈ ⨆ p : toAddSubmonoid '' s, (p : AddSubmonoid M)) ?_ ?_ (fun x y _ _ ↦ ?_) h · rintro ⟨-, ⟨p : Submodule R M, hp : p ∈ s, rfl⟩⟩ x (hx : x ∈ p) diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index aaa15d8a6209e..3fdf3e370cb38 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -10,32 +10,12 @@ import Mathlib.Order.Defs /-! # `NeZero` typeclass -We create a typeclass `NeZero n` which carries around the fact that `(n : R) ≠ 0`. +We give basic facts about the `NeZero n` typeclass. -## Main declarations - -* `NeZero`: `n ≠ 0` as a typeclass. -/ variable {R : Type*} [Zero R] -/-- A type-class version of `n ≠ 0`. -/ -class NeZero (n : R) : Prop where - /-- The proposition that `n` is not zero. -/ - out : n ≠ 0 - -theorem NeZero.ne (n : R) [h : NeZero n] : n ≠ 0 := - h.out - -theorem NeZero.ne' (n : R) [h : NeZero n] : 0 ≠ n := - h.out.symm - -theorem neZero_iff {n : R} : NeZero n ↔ n ≠ 0 := - ⟨fun h ↦ h.out, NeZero.mk⟩ - -@[simp] lemma neZero_zero_iff_false {α : Type*} [Zero α] : NeZero (0 : α) ↔ False := - ⟨fun h ↦ h.ne rfl, fun h ↦ h.elim⟩ - theorem not_neZero {n : R} : ¬NeZero n ↔ n = 0 := by simp [neZero_iff] theorem eq_zero_or_neZero (a : R) : a = 0 ∨ NeZero a := @@ -77,10 +57,6 @@ namespace NeZero variable {M : Type*} {x : M} -instance succ {n : ℕ} : NeZero (n + 1) := ⟨n.succ_ne_zero⟩ - theorem of_pos [Preorder M] [Zero M] (h : 0 < x) : NeZero x := ⟨ne_of_gt h⟩ end NeZero - -lemma Nat.pos_of_neZero (n : ℕ) [NeZero n] : 0 < n := Nat.pos_of_ne_zero (NeZero.ne _) diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 3d5a671d2a7f8..900c60857b3b1 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -164,7 +164,7 @@ variable [CanonicallyOrderedCommMonoid M] {l : List M} @[to_additive] lemma prod_eq_one_iff : l.prod = 1 ↔ ∀ x ∈ l, x = (1 : M) := ⟨all_one_of_le_one_le_of_prod_eq_one fun _ _ => one_le _, fun h => by - rw [List.eq_replicate.2 ⟨_, h⟩, prod_replicate, one_pow] + rw [List.eq_replicate_iff.2 ⟨_, h⟩, prod_replicate, one_pow] · exact (length l) · rfl⟩ diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 18797cc6170d3..6bd51a8b46a96 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -222,13 +222,21 @@ theorem floor_eq_zero : ⌊a⌋₊ = 0 ↔ a < 1 := by rw [← lt_one_iff, ← @cast_one α] exact floor_lt' Nat.one_ne_zero +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem floor_eq_iff (ha : 0 ≤ a) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff ha, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt ha, Nat.lt_add_one_iff, - le_antisymm_iff, and_comm] + le_antisymm_iff, _root_.and_comm] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem floor_eq_iff' (hn : n ≠ 0) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff' hn, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt' (Nat.add_one_ne_zero n), - Nat.lt_add_one_iff, le_antisymm_iff, and_comm] + Nat.lt_add_one_iff, le_antisymm_iff, _root_.and_comm] theorem floor_eq_on_Ico (n : ℕ) : ∀ a ∈ (Set.Ico n (n + 1) : Set α), ⌊a⌋₊ = n := fun _ ⟨h₀, h₁⟩ => (floor_eq_iff <| n.cast_nonneg.trans h₀).mpr ⟨h₀, h₁⟩ @@ -320,10 +328,14 @@ theorem floor_lt_ceil_of_lt_of_pos {a b : α} (h : a < b) (h' : 0 < b) : ⌊a⌋ exact h.trans_le (le_ceil _) · rwa [floor_of_nonpos ha.le, lt_ceil, Nat.cast_zero] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem ceil_eq_iff (hn : n ≠ 0) : ⌈a⌉₊ = n ↔ ↑(n - 1) < a ∧ a ≤ n := by rw [← ceil_le, ← not_le, ← ceil_le, not_le, tsub_lt_iff_right (Nat.add_one_le_iff.2 (pos_iff_ne_zero.2 hn)), Nat.lt_add_one_iff, - le_antisymm_iff, and_comm] + le_antisymm_iff, _root_.and_comm] @[simp] theorem preimage_ceil_zero : (Nat.ceil : α → ℕ) ⁻¹' {0} = Iic 0 := @@ -1054,7 +1066,7 @@ theorem fract_div_intCast_eq_div_intCast_mod {m : ℤ} {n : ℕ} : obtain ⟨l₀, rfl | rfl⟩ := l.eq_nat_or_neg · rw [cast_natCast, ← natCast_mod, cast_natCast, fract_div_natCast_eq_div_natCast_mod] · rw [Right.nonneg_neg_iff, natCast_nonpos_iff] at hl - simp [hl, zero_mod] + simp [hl] obtain ⟨m₀, rfl | rfl⟩ := m.eq_nat_or_neg · exact this (ofNat_nonneg m₀) let q := ⌈↑m₀ / (n : k)⌉ @@ -1694,4 +1706,4 @@ def evalIntCeil : PositivityExt where eval {u α} _zα _pα e := do end Mathlib.Meta.Positivity -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1800 diff --git a/Mathlib/Algebra/Order/Group/Abs.lean b/Mathlib/Algebra/Order/Group/Abs.lean index 3d6763c037cfa..9693c189cfb6c 100644 --- a/Mathlib/Algebra/Order/Group/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Abs.lean @@ -94,7 +94,7 @@ theorem apply_abs_le_mul_of_one_le {β : Type*} [MulOneClass β] [Preorder β] theorem abs_add (a b : α) : |a + b| ≤ |a| + |b| := abs_le.2 ⟨(neg_add |a| |b|).symm ▸ - add_le_add ((@neg_le α ..).2 <| neg_le_abs _) ((@neg_le α ..).2 <| neg_le_abs _), + add_le_add (neg_le.2 <| neg_le_abs _) (neg_le.2 <| neg_le_abs _), add_le_add (le_abs_self _) (le_abs_self _)⟩ theorem abs_add' (a b : α) : |a| ≤ |b| + |b + a| := by simpa using abs_add (-b) (b + a) @@ -122,7 +122,7 @@ theorem sub_lt_of_abs_sub_lt_right (h : |a - b| < c) : a - c < b := sub_lt_of_abs_sub_lt_left (abs_sub_comm a b ▸ h) theorem abs_sub_abs_le_abs_sub (a b : α) : |a| - |b| ≤ |a - b| := - (@sub_le_iff_le_add α ..).2 <| + sub_le_iff_le_add.2 <| calc |a| = |a - b + b| := by rw [sub_add_cancel] _ ≤ |a - b| + |b| := abs_add _ _ diff --git a/Mathlib/Algebra/Order/Group/Cone.lean b/Mathlib/Algebra/Order/Group/Cone.lean index a49307108ac8e..59f273599f45a 100644 --- a/Mathlib/Algebra/Order/Group/Cone.lean +++ b/Mathlib/Algebra/Order/Group/Cone.lean @@ -18,8 +18,8 @@ cones in groups and the corresponding ordered groups. -/ /-- `AddGroupConeClass S G` says that `S` is a type of cones in `G`. -/ -class AddGroupConeClass (S G : Type*) [AddCommGroup G] [SetLike S G] extends - AddSubmonoidClass S G : Prop where +class AddGroupConeClass (S : Type*) (G : outParam Type*) [AddCommGroup G] [SetLike S G] + extends AddSubmonoidClass S G : Prop where eq_zero_of_mem_of_neg_mem {C : S} {a : G} : a ∈ C → -a ∈ C → a = 0 /-- `GroupConeClass S G` says that `S` is a type of cones in `G`. -/ diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index a47229cb3489f..53b53337c45d3 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -21,9 +21,9 @@ The reason is that we did not want to change existing names in the library. -/ /- -`NeZero` should not be needed at this point in the ordered algebraic hierarchy. +`NeZero` theory should not be needed at this point in the ordered algebraic hierarchy. -/ -assert_not_exists NeZero +assert_not_imported Mathlib.Algebra.NeZero open Function @@ -175,13 +175,11 @@ variable [OrderedCommGroup α] {a b : α} @[to_additive (attr := gcongr) neg_le_neg] theorem inv_le_inv' : a ≤ b → b⁻¹ ≤ a⁻¹ := - -- Porting note: explicit type annotation was not needed before. - (@inv_le_inv_iff α ..).mpr + inv_le_inv_iff.mpr @[to_additive (attr := gcongr) neg_lt_neg] theorem inv_lt_inv' : a < b → b⁻¹ < a⁻¹ := - -- Porting note: explicit type annotation was not needed before. - (@inv_lt_inv_iff α ..).mpr + inv_lt_inv_iff.mpr -- The additive version is also a `linarith` lemma. @[to_additive] diff --git a/Mathlib/Algebra/Order/Hom/Basic.lean b/Mathlib/Algebra/Order/Hom/Basic.lean index dbcb880887ff3..b97b19b03d22d 100644 --- a/Mathlib/Algebra/Order/Hom/Basic.lean +++ b/Mathlib/Algebra/Order/Hom/Basic.lean @@ -73,29 +73,33 @@ variable {ι F α β γ δ : Type*} /-! ### Basics -/ /-- `NonnegHomClass F α β` states that `F` is a type of nonnegative morphisms. -/ -class NonnegHomClass (F α β : Type*) [Zero β] [LE β] [FunLike F α β] : Prop where +class NonnegHomClass (F : Type*) (α β : outParam Type*) [Zero β] [LE β] [FunLike F α β] : Prop where /-- the image of any element is non negative. -/ apply_nonneg (f : F) : ∀ a, 0 ≤ f a /-- `SubadditiveHomClass F α β` states that `F` is a type of subadditive morphisms. -/ -class SubadditiveHomClass (F α β : Type*) [Add α] [Add β] [LE β] [FunLike F α β] : Prop where +class SubadditiveHomClass (F : Type*) (α β : outParam Type*) + [Add α] [Add β] [LE β] [FunLike F α β] : Prop where /-- the image of a sum is less or equal than the sum of the images. -/ map_add_le_add (f : F) : ∀ a b, f (a + b) ≤ f a + f b /-- `SubmultiplicativeHomClass F α β` states that `F` is a type of submultiplicative morphisms. -/ @[to_additive SubadditiveHomClass] -class SubmultiplicativeHomClass (F α β : Type*) [Mul α] [Mul β] [LE β] [FunLike F α β] : Prop where +class SubmultiplicativeHomClass (F : Type*) (α β : outParam (Type*)) [Mul α] [Mul β] [LE β] + [FunLike F α β] : Prop where /-- the image of a product is less or equal than the product of the images. -/ map_mul_le_mul (f : F) : ∀ a b, f (a * b) ≤ f a * f b /-- `MulLEAddHomClass F α β` states that `F` is a type of subadditive morphisms. -/ @[to_additive SubadditiveHomClass] -class MulLEAddHomClass (F α β : Type*) [Mul α] [Add β] [LE β] [FunLike F α β] : Prop where +class MulLEAddHomClass (F : Type*) (α β : outParam Type*) [Mul α] [Add β] [LE β] [FunLike F α β] : + Prop where /-- the image of a product is less or equal than the sum of the images. -/ map_mul_le_add (f : F) : ∀ a b, f (a * b) ≤ f a + f b /-- `NonarchimedeanHomClass F α β` states that `F` is a type of non-archimedean morphisms. -/ -class NonarchimedeanHomClass (F α β : Type*) [Add α] [LinearOrder β] [FunLike F α β] : Prop where +class NonarchimedeanHomClass (F : Type*) (α β : outParam Type*) + [Add α] [LinearOrder β] [FunLike F α β] : Prop where /-- the image of a sum is less or equal than the maximum of the images. -/ map_add_le_max (f : F) : ∀ a b, f (a + b) ≤ max (f a) (f b) @@ -154,7 +158,8 @@ end Mathlib.Meta.Positivity group `α`. You should extend this class when you extend `AddGroupSeminorm`. -/ -class AddGroupSeminormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] +class AddGroupSeminormClass (F : Type*) (α β : outParam Type*) + [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] extends SubadditiveHomClass F α β : Prop where /-- The image of zero is zero. -/ map_zero (f : F) : f 0 = 0 @@ -165,7 +170,8 @@ class AddGroupSeminormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoi You should extend this class when you extend `GroupSeminorm`. -/ @[to_additive] -class GroupSeminormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] [FunLike F α β] +class GroupSeminormClass (F : Type*) (α β : outParam Type*) + [Group α] [OrderedAddCommMonoid β] [FunLike F α β] extends MulLEAddHomClass F α β : Prop where /-- The image of one is zero. -/ map_one_eq_zero (f : F) : f 1 = 0 @@ -176,7 +182,8 @@ class GroupSeminormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] `α`. You should extend this class when you extend `AddGroupNorm`. -/ -class AddGroupNormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] +class AddGroupNormClass (F : Type*) (α β : outParam Type*) + [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] extends AddGroupSeminormClass F α β : Prop where /-- The argument is zero if its image under the map is zero. -/ eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0 @@ -185,7 +192,8 @@ class AddGroupNormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β You should extend this class when you extend `GroupNorm`. -/ @[to_additive] -class GroupNormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] [FunLike F α β] +class GroupNormClass (F : Type*) (α β : outParam Type*) + [Group α] [OrderedAddCommMonoid β] [FunLike F α β] extends GroupSeminormClass F α β : Prop where /-- The argument is one if its image under the map is zero. -/ eq_one_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 1 @@ -275,20 +283,23 @@ theorem map_pos_of_ne_one [Group α] [LinearOrderedAddCommMonoid β] [GroupNormC /-- `RingSeminormClass F α` states that `F` is a type of `β`-valued seminorms on the ring `α`. You should extend this class when you extend `RingSeminorm`. -/ -class RingSeminormClass (F α β : Type*) [NonUnitalNonAssocRing α] [OrderedSemiring β] - [FunLike F α β] extends AddGroupSeminormClass F α β, SubmultiplicativeHomClass F α β : Prop +class RingSeminormClass (F : Type*) (α β : outParam Type*) + [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] + extends AddGroupSeminormClass F α β, SubmultiplicativeHomClass F α β : Prop /-- `RingNormClass F α` states that `F` is a type of `β`-valued norms on the ring `α`. You should extend this class when you extend `RingNorm`. -/ -class RingNormClass (F α β : Type*) [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class RingNormClass (F : Type*) (α β : outParam Type*) + [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends RingSeminormClass F α β, AddGroupNormClass F α β : Prop /-- `MulRingSeminormClass F α` states that `F` is a type of `β`-valued multiplicative seminorms on the ring `α`. You should extend this class when you extend `MulRingSeminorm`. -/ -class MulRingSeminormClass (F α β : Type*) [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class MulRingSeminormClass (F : Type*) (α β : outParam Type*) + [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends AddGroupSeminormClass F α β, MonoidWithZeroHomClass F α β : Prop -- Lower the priority of these instances since they require synthesizing an order structure. @@ -299,7 +310,8 @@ attribute [instance 50] ring `α`. You should extend this class when you extend `MulRingNorm`. -/ -class MulRingNormClass (F α β : Type*) [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class MulRingNormClass (F : Type*) (α β : outParam Type*) + [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends MulRingSeminormClass F α β, AddGroupNormClass F α β : Prop -- See note [out-param inheritance] diff --git a/Mathlib/Algebra/Order/Monoid/Prod.lean b/Mathlib/Algebra/Order/Monoid/Prod.lean index c6998f304353c..e3d9ab5369849 100644 --- a/Mathlib/Algebra/Order/Monoid/Prod.lean +++ b/Mathlib/Algebra/Order/Monoid/Prod.lean @@ -40,7 +40,7 @@ instance [CanonicallyOrderedCommMonoid α] [CanonicallyOrderedCommMonoid β] : CanonicallyOrderedCommMonoid (α × β) := { (inferInstance : OrderedCommMonoid _), (inferInstance : OrderBot _), (inferInstance : ExistsMulOfLE _) with - le_self_mul := fun _ _ ↦ ⟨le_self_mul, le_self_mul⟩ } + le_self_mul := fun _ _ ↦ le_def.mpr ⟨le_self_mul, le_self_mul⟩ } namespace Lex diff --git a/Mathlib/Algebra/Order/Ring/Cone.lean b/Mathlib/Algebra/Order/Ring/Cone.lean index 4ddd18e2f67d0..71c21076f0093 100644 --- a/Mathlib/Algebra/Order/Ring/Cone.lean +++ b/Mathlib/Algebra/Order/Ring/Cone.lean @@ -19,7 +19,7 @@ cones in rings and the corresponding ordered rings. -/ /-- `RingConeClass S R` says that `S` is a type of cones in `R`. -/ -class RingConeClass (S R : Type*) [Ring R] [SetLike S R] +class RingConeClass (S : Type*) (R : outParam Type*) [Ring R] [SetLike S R] extends AddGroupConeClass S R, SubsemiringClass S R : Prop /-- A (positive) cone in a ring is a subsemiring that diff --git a/Mathlib/Algebra/Order/ZeroLEOne.lean b/Mathlib/Algebra/Order/ZeroLEOne.lean index ecd0e356d73ad..538c8f6985b42 100644 --- a/Mathlib/Algebra/Order/ZeroLEOne.lean +++ b/Mathlib/Algebra/Order/ZeroLEOne.lean @@ -4,7 +4,6 @@ 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.Order.Basic -import Mathlib.Algebra.NeZero /-! # Typeclass expressing `0 ≤ 1`. diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index 32db3b30a3a97..cef12c507490a 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -68,7 +68,7 @@ theorem degree_list_sum_le (l : List S[X]) : degree l.sum ≤ (l.map natDegree). rw [← List.foldr_max_of_ne_nil] · congr contrapose! h - rw [List.map_eq_nil] at h + rw [List.map_eq_nil_iff] at h simp [h] theorem natDegree_list_prod_le (l : List S[X]) : natDegree l.prod ≤ (l.map natDegree).sum := by diff --git a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean index 1a1dd67ce6467..01b1c8ec0561e 100644 --- a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean @@ -50,7 +50,7 @@ noncomputable def cardPowDegree : AbsoluteValue Fq[X] ℤ := · rfl exact pow_nonneg (Int.ofNat_zero_le _) _ eq_zero' := fun p => - ite_eq_left_iff.trans <| + ite_eq_left_iff.trans ⟨fun h => by contrapose! h exact ⟨h, (pow_pos _).ne'⟩, absurd⟩ diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index 25857d52638e5..37a70431efb16 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -61,8 +61,8 @@ attribute [nolint docBlame] CentroidHom.toAddMonoidHom /-- `CentroidHomClass F α` states that `F` is a type of centroid homomorphisms. You should extend this class when you extend `CentroidHom`. -/ -class CentroidHomClass (F α : Type*) [NonUnitalNonAssocSemiring α] [FunLike F α α] extends - AddMonoidHomClass F α α : Prop where +class CentroidHomClass (F : Type*) (α : outParam Type*) + [NonUnitalNonAssocSemiring α] [FunLike F α α] extends AddMonoidHomClass F α α : Prop where /-- Commutativity of centroid homomorphims with left multiplication. -/ map_mul_left (f : F) (a b : α) : f (a * b) = a * f b /-- Commutativity of centroid homomorphims with right multiplication. -/ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 8d540524abfb4..088d5ce242d13 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -231,10 +231,14 @@ lemma even_or_odd (n : ℕ) : Even n ∨ Odd n := (even_xor_odd n).or lemma even_or_odd' (n : ℕ) : ∃ k, n = 2 * k ∨ n = 2 * k + 1 := by simpa only [← two_mul, exists_or, Odd, Even] using even_or_odd n +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma even_xor_odd' (n : ℕ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by obtain ⟨k, rfl⟩ | ⟨k, rfl⟩ := even_or_odd n <;> use k · simpa only [← two_mul, eq_self_iff_true, xor_true] using (succ_ne_self (2 * k)).symm - · simpa only [xor_true, xor_comm] using (succ_ne_self _) + · simpa only [xor_true, _root_.xor_comm] using (succ_ne_self _) lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m + n) % 2 = 1 := ((even_or_odd m).elim fun hm ↦ by rw [even_iff.1 hm, odd_iff.1 (hm.add_odd hn)]) fun hm ↦ by diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index dece9ec6923b9..4d2d0317d2fcf 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -72,7 +72,7 @@ section SubringClass /-- `SubringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative submonoid and an additive subgroup. -/ -class SubringClass (S : Type*) (R : Type u) [Ring R] [SetLike S R] extends +class SubringClass (S : Type*) (R : outParam (Type u)) [Ring R] [SetLike S R] extends SubsemiringClass S R, NegMemClass S R : Prop -- See note [lower instance priority] diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 2c20870f69901..f775d6a04fdda 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -59,12 +59,12 @@ section SubsemiringClass /-- `SubsemiringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative and an additive submonoid. -/ -class SubsemiringClass (S : Type*) (R : Type u) [NonAssocSemiring R] +class SubsemiringClass (S : Type*) (R : outParam (Type u)) [NonAssocSemiring R] [SetLike S R] extends SubmonoidClass S R, AddSubmonoidClass S R : Prop -- See note [lower instance priority] instance (priority := 100) SubsemiringClass.addSubmonoidWithOneClass (S : Type*) - (R : Type u) [NonAssocSemiring R] [SetLike S R] [h : SubsemiringClass S R] : + (R : Type u) {_ : NonAssocSemiring R} [SetLike S R] [h : SubsemiringClass S R] : AddSubmonoidWithOneClass S R := { h with } diff --git a/Mathlib/Algebra/Star/Free.lean b/Mathlib/Algebra/Star/Free.lean index eefbd4ba0c208..b29ed2137d81e 100644 --- a/Mathlib/Algebra/Star/Free.lean +++ b/Mathlib/Algebra/Star/Free.lean @@ -48,7 +48,7 @@ instance : StarRing (FreeAlgebra R X) where unfold Star.star simp only [Function.comp_apply] let y := lift R (X := X) (MulOpposite.op ∘ ι R) - apply induction (C := fun x ↦ (y (y x).unop).unop = x) _ _ _ _ x + refine induction (C := fun x ↦ (y (y x).unop).unop = x) _ _ ?_ ?_ ?_ ?_ x · intros simp only [AlgHom.commutes, MulOpposite.algebraMap_apply, MulOpposite.unop_op] · intros diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index b81f47d19763c..b3aa7de64c852 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -187,10 +187,11 @@ lemma smul_section_apply (r : R) (U : Opens (PrimeSpectrum.Top R)) lemma smul_stalk_no_nonzero_divisor {x : PrimeSpectrum R} (r : x.asIdeal.primeCompl) (st : (tildeInModuleCat M).stalk x) (hst : r.1 • st = 0) : st = 0 := by - refine Limits.Concrete.colimit_no_zero_smul_divisor (hx := hst) - ⟨op ⟨PrimeSpectrum.basicOpen r.1, r.2⟩, fun U i s hs ↦ Subtype.eq <| funext fun pt ↦ ?_⟩ - exact LocalizedModule.eq_zero_of_smul_eq_zero (hx := congr_fun (Subtype.ext_iff.1 hs) pt) <| - i.unop pt |>.2 + refine Limits.Concrete.colimit_no_zero_smul_divisor + _ _ _ ⟨op ⟨PrimeSpectrum.basicOpen r.1, r.2⟩, fun U i s hs ↦ Subtype.eq <| funext fun pt ↦ ?_⟩ + _ hst + apply LocalizedModule.eq_zero_of_smul_eq_zero _ (i.unop pt).2 _ + (congr_fun (Subtype.ext_iff.1 hs) pt) /-- If `U` is an open subset of `Spec R`, this is the morphism of `R`-modules from `M` to diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index c85901ed45bc9..cf3a6124cc51f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -198,7 +198,15 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch {U : X.Opens} (hU : IsAffineOpen U) (x f : Γ(X, U)) (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by rw [← map_zero (X.presheaf.map (homOfLE <| X.basicOpen_le f : X.basicOpen f ⟶ U).op)] at H - obtain ⟨n, e⟩ := (hU.isLocalization_basicOpen f).exists_of_eq H + #adaptation_note + /-- + Prior to nightly-2024-09-29, we could use dot notation here: + `(hU.isLocalization_basicOpen f).exists_of_eq H` + This is no longer possible; + likely changing the signature of `IsLocalization.Away.exists_of_eq` is in order. + -/ + obtain ⟨n, e⟩ := + @IsLocalization.Away.exists_of_eq _ _ _ _ _ _ (hU.isLocalization_basicOpen f) _ _ H exact ⟨n, by simpa [mul_comm x] using e⟩ /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index c65f7aa086fbc..5ace362109bfb 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -431,7 +431,7 @@ instance pullback_fst_of_right : IsOpenImmersion (pullback.fst g f) := by rw [← pullbackSymmetry_hom_comp_snd] -- Porting note: was just `infer_instance`, it is a bit weird that no explicit class instance is -- provided but still class inference fail to find this - exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ + exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ instance pullback_to_base [IsOpenImmersion g] : IsOpenImmersion (limit.π (cospan f g) WalkingCospan.one) := by @@ -439,14 +439,14 @@ instance pullback_to_base [IsOpenImmersion g] : change IsOpenImmersion (_ ≫ f) -- Porting note: was just `infer_instance`, it is a bit weird that no explicit class instance is -- provided but still class inference fail to find this - exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ + exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ instance forgetToTopPreservesOfLeft : PreservesLimit (cospan f g) Scheme.forgetToTop := by delta Scheme.forgetToTop - apply @Limits.compPreservesLimit (K := cospan f g) (F := forget) + refine @Limits.compPreservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) (G := LocallyRingedSpace.forgetToTop) ?_ ?_ · infer_instance - apply @preservesLimitOfIsoDiagram (F := _) _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ + refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ dsimp [LocallyRingedSpace.forgetToTop] infer_instance diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 139077e8ace13..938d2b0845645 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -294,7 +294,12 @@ theorem mem_carrier_iff_of_mem (hm : 0 < m) (q : Spec.T A⁰_ f) (a : A) {n} (hn trans (HomogeneousLocalization.mk ⟨m * n, ⟨proj 𝒜 n a ^ m, by rw [← smul_eq_mul]; mem_tac⟩, ⟨f ^ n, by rw [mul_comm]; mem_tac⟩, ⟨_, rfl⟩⟩ : A⁰_ f) ∈ q.asIdeal · refine ⟨fun h ↦ h n, fun h i ↦ if hi : i = n then hi ▸ h else ?_⟩ - convert zero_mem q.asIdeal + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 + we need to specify the implicit arguments by `inferInstance`. + -/ + convert @zero_mem _ _ inferInstance inferInstance _ q.asIdeal apply HomogeneousLocalization.val_injective simp only [proj_apply, decompose_of_mem_ne _ hn (Ne.symm hi), zero_pow hm.ne', HomogeneousLocalization.val_mk, Localization.mk_zero, HomogeneousLocalization.val_zero] @@ -806,7 +811,7 @@ If `f ∈ A` is a homogeneous element of positive degree, then the projective sp -/ def projIsoSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (Proj| pbo f) ≅ (Spec (A⁰_ f)) := - @asIso (f := toSpec 𝒜 f) (isIso_toSpec 𝒜 f f_deg hm) + @asIso _ _ _ _ (f := toSpec 𝒜 f) (isIso_toSpec 𝒜 f f_deg hm) /-- This is the scheme `Proj(A)` for any `ℕ`-graded ring `A`. diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 2dcfa434d2017..1b43cf9c18bfd 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -507,7 +507,7 @@ def openCoverOfBase' (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCove pasteVertIsPullback rfl (pullbackIsPullback g (𝒰.map i)) (pullbackIsPullback (pullback.snd g (𝒰.map i)) (pullback.snd f (𝒰.map i))) refine - @openCoverOfIsIso + @openCoverOfIsIso _ _ (f := (pullbackSymmetry _ _).hom ≫ (limit.isoLimitCone ⟨_, this⟩).inv ≫ pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) ?_ ?_) inferInstance · simp [← pullback.condition] @@ -583,7 +583,8 @@ the morphism `Spec (S ⊗[R] T) ⟶ Spec T` obtained by applying `Spec.map` to t -/ @[reassoc (attr := simp)] lemma pullbackSpecIso_inv_snd : - (pullbackSpecIso R S T).inv ≫ pullback.snd _ _ = Spec.map (ofHom (toRingHom includeRight)) := + (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 diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index d9bd84e10d143..49ef36b0ddbd9 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1043,7 +1043,7 @@ def sigmaCompositionAux (a : Composition n) (b : Composition a.length) a.blocks_pos (by rw [← a.blocks.join_splitWrtComposition b] - exact mem_join_of_mem (List.getElem_mem _ _ _) hi) + exact mem_join_of_mem (List.getElem_mem _) hi) blocks_sum := by simp [Composition.blocksFun, getElem_map, Composition.gather] theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length) @@ -1109,7 +1109,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i have : sizeUpTo b i + Nat.succ j = (sizeUpTo b i + j).succ := rfl rw [this, sizeUpTo_succ _ D, IHj A, sizeUpTo_succ _ B] simp only [sigmaCompositionAux, add_assoc, add_left_inj, Fin.val_mk] - rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop, getElem_take _ _ C] + rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop, getElem_take' _ _ C] /-- Natural equivalence between `(Σ (a : Composition n), Composition a.length)` and `(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocksFun i))`, that shows up as a diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index a085109299f03..180ca90ca5288 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -605,7 +605,7 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] lemma cfcHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcHom ha = ha.spectrumRestricts.starAlgHom (cfcHom ha.isStarNormal) (f := Complex.reCLM) := - ha.spectrumRestricts.cfcHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding + ha.spectrumRestricts.cfcHom_eq_restrict _ Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfc_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : @@ -626,7 +626,7 @@ variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module lemma cfcₙHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcₙHom ha = (ha.quasispectrumRestricts.2).nonUnitalStarAlgHom (cfcₙHom ha.isStarNormal) (f := Complex.reCLM) := - ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding + ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict _ Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfcₙ_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : @@ -650,12 +650,12 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [PartialOrder A] [StarRing A] lemma cfcHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcHom ha = (SpectrumRestricts.nnreal_of_nonneg ha).starAlgHom (cfcHom (IsSelfAdjoint.of_nonneg ha)) := by - apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict isUniformEmbedding_subtype_val + apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict _ isUniformEmbedding_subtype_val lemma cfc_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfc f a = cfc (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam - apply (SpectrumRestricts.nnreal_of_nonneg ha).cfc_eq_restrict + apply (SpectrumRestricts.nnreal_of_nonneg ha).cfc_eq_restrict _ isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqReal @@ -674,13 +674,13 @@ variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [PartialOrder A] [St lemma cfcₙHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcₙHom ha = (QuasispectrumRestricts.nnreal_of_nonneg ha).nonUnitalStarAlgHom (cfcₙHom (IsSelfAdjoint.of_nonneg ha)) := by - apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙHom_eq_restrict + apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙHom_eq_restrict _ isUniformEmbedding_subtype_val lemma cfcₙ_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfcₙ f a = cfcₙ (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam - apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙ_eq_restrict + apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙ_eq_restrict _ isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqRealNonUnital diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 3bcd0795b06d8..a298ba04ac176 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -421,7 +421,7 @@ lemma cfcₙ_comp (g f : R → R) (a : A) ext simp rw [cfcₙ_apply .., cfcₙ_apply f a, - cfcₙ_apply _ (by convert hg) (ha := cfcₙHom_predicate (show p a from ha) _) , + cfcₙ_apply _ _ (by convert hg) (ha := cfcₙHom_predicate (show p a from ha) _), ← cfcₙHom_comp _ _] swap · exact ⟨.mk _ <| hf.restrict.codRestrict fun x ↦ by rw [sp_eq]; use x.1; simp, Subtype.ext hf0⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 89017072beb7c..8b6473cba0de3 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -723,7 +723,7 @@ noncomputable def cfcUnits (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) lemma cfcUnits_pow (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) (n : ℕ) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : (cfcUnits f a hf') ^ n = - cfcUnits (forall₂_imp (fun _ _ ↦ pow_ne_zero n) hf') (hf := hf.pow n) := by + cfcUnits _ _ (forall₂_imp (fun _ _ ↦ pow_ne_zero n) hf') (hf := hf.pow n) := by ext cases n with | zero => simp [cfc_const_one R a] @@ -778,7 +778,7 @@ lemma cfcUnits_zpow (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) (n : ℤ) | negSucc n => simp only [zpow_negSucc, ← inv_pow] ext - exact cfc_pow (hf := hf.inv₀ hf') _ |>.symm + exact cfc_pow (hf := hf.inv₀ hf') .. |>.symm lemma cfc_zpow (a : Aˣ) (n : ℤ) (ha : p a := by cfc_tac) : cfc (fun x : R ↦ x ^ n) (a : A) = ↑(a ^ n) := by diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index edda887e12373..5ed888191d0e0 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -416,7 +416,7 @@ theorem norm_deriv_eq_norm_fderiv : ‖deriv f x‖ = ‖fderiv 𝕜 f x‖ := b theorem DifferentiableAt.derivWithin (h : DifferentiableAt 𝕜 f x) (hxs : UniqueDiffWithinAt 𝕜 s x) : derivWithin f s x = deriv f x := by - unfold derivWithin deriv + unfold _root_.derivWithin deriv rw [h.fderivWithin hxs] theorem HasDerivWithinAt.deriv_eq_zero (hd : HasDerivWithinAt f 0 s x) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index d867ef086d9c7..990bb02bd9599 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -573,10 +573,10 @@ theorem hasStrictFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasStrictFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasStrictFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj l.attach[i.cast l.length_attach.symm]) + smulRight (proj l.attach[i.cast List.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod'.congr_fderiv <| Eq.symm <| - Finset.sum_equiv (finCongr l.length_attach.symm) (by simp) (by simp) + Finset.sum_equiv (finCongr List.length_attach.symm) (by simp) (by simp) @[fun_prop] theorem hasFDerivAt_list_prod' [Fintype ι] {l : List ι} {x : ι → 𝔸'} : @@ -596,7 +596,7 @@ theorem hasFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj l.attach[i.cast l.length_attach.symm]) + smulRight (proj l.attach[i.cast List.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod_attach'.hasFDerivAt @@ -648,7 +648,7 @@ theorem HasStrictFDerivAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasStrictFDerivAt_list_prod_finRange'.comp x - (hasStrictFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem ..))) ?_ + (hasStrictFDerivAt_pi.mpr fun i ↦ h l[i] (List.getElem_mem ..))) ?_ ext m simp [← List.map_map] @@ -663,7 +663,7 @@ theorem HasFDerivAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasFDerivAt_list_prod_finRange'.comp x - (hasFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem i i.isLt))) ?_ + (hasFDerivAt_pi.mpr fun i ↦ h l[i] (List.getElem_mem i.isLt))) ?_ ext m simp [← List.map_map] diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index b29417e75a1af..e2e23c0b46a9d 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -101,7 +101,7 @@ variable {E} theorem _root_.Continuous.inner_ {f g : ℝ → E} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => inner_ 𝕜 (f x) (g x) := by - unfold inner_ + unfold _root_.inner_ fun_prop theorem inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := by diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index c334bfd9e636b..770630be496f3 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -747,7 +747,7 @@ theorem norm_mkPiAlgebraFin_succ_le : ‖ContinuousMultilinearMap.mkPiAlgebraFin simp only [ContinuousMultilinearMap.mkPiAlgebraFin_apply, one_mul, List.ofFn_eq_map, Fin.prod_univ_def, Multiset.map_coe, Multiset.prod_coe] refine (List.norm_prod_le' ?_).trans_eq ?_ - · rw [Ne, List.map_eq_nil, List.finRange_eq_nil] + · rw [Ne, List.map_eq_nil_iff, List.finRange_eq_nil] exact Nat.succ_ne_zero _ rw [List.map_map, Function.comp_def] diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 506735693822c..31f51791cdb31 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -33,7 +33,7 @@ integrate, integration, integrable, integrability -/ -open Real Nat Set Finset +open Real Set Finset open scoped Real Interval @@ -580,8 +580,9 @@ theorem integral_mul_rpow_one_add_sq {t : ℝ} (ht : t ≠ -1) : end RpowCpow -/-! ### Integral of `sin x ^ n` -/ +open Nat +/-! ### Integral of `sin x ^ n` -/ theorem integral_sin_pow_aux : (∫ x in a..b, sin x ^ (n + 2)) = diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index dfd9f1a9f4692..6408603738450 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -517,7 +517,7 @@ theorem cos_eq_one_iff (x : ℝ) : cos x = 1 ↔ ∃ n : ℤ, (n : ℝ) * (2 * (Int.emod_two_eq_zero_or_one n).elim (fun hn0 => by rwa [← mul_assoc, ← @Int.cast_two ℝ, ← Int.cast_mul, - Int.ediv_mul_cancel ((Int.dvd_iff_emod_eq_zero _ _).2 hn0)]) + Int.ediv_mul_cancel (Int.dvd_iff_emod_eq_zero.2 hn0)]) fun hn1 => by rw [← Int.emod_add_ediv n 2, hn1, Int.cast_add, Int.cast_one, add_mul, one_mul, add_comm, mul_comm (2 : ℤ), Int.cast_mul, mul_assoc, Int.cast_two] at hn diff --git a/Mathlib/CategoryTheory/ConnectedComponents.lean b/Mathlib/CategoryTheory/ConnectedComponents.lean index cc86ae90a2167..0d205fcbb0170 100644 --- a/Mathlib/CategoryTheory/ConnectedComponents.lean +++ b/Mathlib/CategoryTheory/ConnectedComponents.lean @@ -7,7 +7,6 @@ import Mathlib.Data.List.Chain import Mathlib.CategoryTheory.IsConnected import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.FullSubcategory -import Mathlib.Data.List.Infix /-! # Connected components of a category diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 887e41f0debb8..94619f08a21bb 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -527,7 +527,7 @@ instance FinitaryPreExtensive.hasPullbacks_of_inclusions [FinitaryPreExtensive C {α : Type*} (f : X ⟶ Z) {Y : (a : α) → C} (i : (a : α) → Y a ⟶ Z) [Finite α] [hi : IsIso (Sigma.desc i)] (a : α) : HasPullback f (i a) := by apply FinitaryPreExtensive.hasPullbacks_of_is_coproduct (c := Cofan.mk Z i) - exact @IsColimit.ofPointIso (t := Cofan.mk Z i) (P := _) hi + exact @IsColimit.ofPointIso (t := Cofan.mk Z i) (P := _) (i := hi) lemma FinitaryPreExtensive.sigma_desc_iso [FinitaryPreExtensive C] {α : Type} [Finite α] {X : C} {Z : α → C} (π : (a : α) → Z a ⟶ X) {Y : C} (f : Y ⟶ X) (hπ : IsIso (Sigma.desc π)) : @@ -536,7 +536,7 @@ lemma FinitaryPreExtensive.sigma_desc_iso [FinitaryPreExtensive C] {α : Type} [ change IsIso (this.coconePointUniqueUpToIso (getColimitCocone _).2).inv infer_instance let this : IsColimit (Cofan.mk X π) := by - refine @IsColimit.ofPointIso (t := Cofan.mk X π) (P := coproductIsCoproduct Z) ?_ + refine @IsColimit.ofPointIso (t := Cofan.mk X π) (P := coproductIsCoproduct Z) (i := ?_) convert hπ simp [coproductIsCoproduct] refine (FinitaryPreExtensive.isUniversal_finiteCoproducts this diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean index 57c9c1c9353e0..fe8433c905fb7 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean @@ -168,7 +168,7 @@ precomposition by `L`. -/ noncomputable def ranAdjunction : (whiskeringLeft C D H).obj L ⊣ L.ran := Adjunction.mkOfHomEquiv { homEquiv := fun F G => - (homEquivOfIsRightKanExtension (α := L.ranCounit.app G) F).symm + (homEquivOfIsRightKanExtension (α := L.ranCounit.app G) _ F).symm homEquiv_naturality_right := fun {F G₁ G₂} β f ↦ hom_ext_of_isRightKanExtension _ (L.ranCounit.app G₂) _ _ (by ext X diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 917484d456465..b005a034c6b49 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -325,7 +325,7 @@ lemma left_tensor_tensorObj₃_ext {j : I} {A : C} (Z : C) (_ ◁ ιTensorObj₃ X₁ X₂ X₃ i₁ i₂ i₃ j h) ≫ f = (_ ◁ ιTensorObj₃ X₁ X₂ X₃ i₁ i₂ i₃ j h) ≫ g) : f = g := by refine (@isColimitOfPreserves C _ C _ _ _ _ ((curriedTensor C).obj Z) _ - (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (H := H) j) hZ).hom_ext ?_ + (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (H := H) (j := j)) hZ).hom_ext ?_ intro ⟨⟨i₁, i₂, i₃⟩, hi⟩ exact h _ _ _ hi diff --git a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean index 034992348ba84..6b3ab2bd33ac2 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean @@ -41,7 +41,7 @@ class HasLiftingProperty : Prop where sq_hasLift : ∀ {f : A ⟶ X} {g : B ⟶ Y} (sq : CommSq f i p g), sq.HasLift instance (priority := 100) sq_hasLift_of_hasLiftingProperty {f : A ⟶ X} {g : B ⟶ Y} - (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := by apply hip.sq_hasLift + (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := hip.sq_hasLift _ namespace HasLiftingProperty diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 22aa87a4a18cd..10e2d68009de0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -81,8 +81,6 @@ attribute [reassoc (attr := simp)] MonoFactorisation.fac attribute [instance] MonoFactorisation.m_mono -attribute [instance] MonoFactorisation.m_mono - namespace MonoFactorisation /-- The obvious factorisation of a monomorphism through itself. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index e22ae9d2d8bed..423fed53f4080 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -280,22 +280,22 @@ noncomputable def isoIsPullback (h : IsPullback fst snd f g) (h' : IsPullback fs @[reassoc (attr := simp)] theorem isoIsPullback_hom_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').hom ≫ fst' = fst := + (h.isoIsPullback _ _ h').hom ≫ fst' = fst := IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.left @[reassoc (attr := simp)] theorem isoIsPullback_hom_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').hom ≫ snd' = snd := + (h.isoIsPullback _ _ h').hom ≫ snd' = snd := IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.right @[reassoc (attr := simp)] theorem isoIsPullback_inv_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').inv ≫ fst = fst' := by + (h.isoIsPullback _ _ h').inv ≫ fst = fst' := by simp only [Iso.inv_comp_eq, isoIsPullback_hom_fst] @[reassoc (attr := simp)] theorem isoIsPullback_inv_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').inv ≫ snd = snd' := by + (h.isoIsPullback _ _ h').inv ≫ snd = snd' := by simp only [Iso.inv_comp_eq, isoIsPullback_hom_snd] end @@ -468,22 +468,22 @@ noncomputable def isoIsPushout (h : IsPushout f g inl inr) (h' : IsPushout f g i @[reassoc (attr := simp)] theorem inl_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inl ≫ (h.isoIsPushout h').hom = inl' := + inl ≫ (h.isoIsPushout _ _ h').hom = inl' := IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.left @[reassoc (attr := simp)] theorem inr_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inr ≫ (h.isoIsPushout h').hom = inr' := + inr ≫ (h.isoIsPushout _ _ h').hom = inr' := IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.right @[reassoc (attr := simp)] theorem inl_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inl' ≫ (h.isoIsPushout h').inv = inl := by + inl' ≫ (h.isoIsPushout _ _ h').inv = inl := by simp only [Iso.comp_inv_eq, inl_isoIsPushout_hom] @[reassoc (attr := simp)] theorem inr_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inr' ≫ (h.isoIsPushout h').inv = inr := by + inr' ≫ (h.isoIsPushout _ _ h').inv = inr := by simp only [Iso.comp_inv_eq, inr_isoIsPushout_hom] end diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index e9d8b8fb9c0ff..0a3a382da8016 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -162,13 +162,13 @@ lemma RespectsIso.postcomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} [IsIso e] (f : X ⟶ Y) (hf : P f) : P (f ≫ e) := RespectsRight.postcomp (Q := isomorphisms C) e ‹IsIso e› f hf -instance RespectsIso.op (P : MorphismProperty C) [h : RespectsIso P] : RespectsIso P.op where - precomp e (_ : IsIso e) f hf := h.postcomp e.unop f.unop hf - postcomp e (_ : IsIso e) f hf := h.precomp e.unop f.unop hf +instance RespectsIso.op (P : MorphismProperty C) [RespectsIso P] : RespectsIso P.op where + precomp e (_ : IsIso e) f hf := postcomp P e.unop f.unop hf + postcomp e (_ : IsIso e) f hf := precomp P e.unop f.unop hf -instance RespectsIso.unop (P : MorphismProperty Cᵒᵖ) [h : RespectsIso P] : RespectsIso P.unop where - precomp e (_ : IsIso e) f hf := h.postcomp e.op f.op hf - postcomp e (_ : IsIso e) f hf := h.precomp e.op f.op hf +instance RespectsIso.unop (P : MorphismProperty Cᵒᵖ) [RespectsIso P] : RespectsIso P.unop where + precomp e (_ : IsIso e) f hf := postcomp P e.op f.op hf + postcomp e (_ : IsIso e) f hf := precomp P e.op f.op hf /-- The closure by isomorphisms of a `MorphismProperty` -/ def isoClosure (P : MorphismProperty C) : MorphismProperty C := @@ -190,11 +190,11 @@ lemma monotone_isoClosure : Monotone (isoClosure (C := C)) := by theorem cancel_left_of_respectsIso (P : MorphismProperty C) [hP : RespectsIso P] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : P (f ≫ g) ↔ P g := - ⟨fun h => by simpa using hP.precomp (inv f) (f ≫ g) h, hP.precomp f g⟩ + ⟨fun h => by simpa using RespectsIso.precomp P (inv f) (f ≫ g) h, RespectsIso.precomp P f g⟩ theorem cancel_right_of_respectsIso (P : MorphismProperty C) [hP : RespectsIso P] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : P (f ≫ g) ↔ P f := - ⟨fun h => by simpa using hP.postcomp (inv g) (f ≫ g) h, hP.postcomp g f⟩ + ⟨fun h => by simpa using RespectsIso.postcomp P (inv g) (f ≫ g) h, RespectsIso.postcomp P g f⟩ theorem arrow_iso_iff (P : MorphismProperty C) [RespectsIso P] {f g : Arrow C} (e : f ≅ g) : P f.hom ↔ P g.hom := by diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index 7f2817071b11a..a20deb1ceff5d 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -85,7 +85,15 @@ lemma eq_induced : haveI := F.reflects_precoherent instance : haveI := F.reflects_precoherent; F.IsDenseSubsite (coherentTopology C) (coherentTopology D) where - functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl + functorPushforward_mem_iff := by + rw [eq_induced F] + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl lemma coverPreserving : haveI := F.reflects_precoherent CoverPreserving (coherentTopology _) (coherentTopology _) F := @@ -181,7 +189,15 @@ lemma eq_induced : haveI := F.reflects_preregular instance : haveI := F.reflects_preregular; F.IsDenseSubsite (regularTopology C) (regularTopology D) where - functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl + functorPushforward_mem_iff := by + rw [eq_induced F] + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl lemma coverPreserving : haveI := F.reflects_preregular CoverPreserving (regularTopology _) (regularTopology _) F := diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 3e453496f9a9d..0d3552925cc54 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -263,7 +263,13 @@ instance : InfSet (GrothendieckTopology C) where /-- See -/ theorem isGLB_sInf (s : Set (GrothendieckTopology C)) : IsGLB s (sInf s) := by refine @IsGLB.of_image _ _ _ _ sieves ?_ _ _ ?_ - · rfl + · #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl · exact _root_.isGLB_sInf _ /-- Construct a complete lattice from the `Inf`, but make the trivial and discrete topologies diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index d78ec3c589a8f..9507a7937e954 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -310,6 +310,6 @@ lemma isTriangulated_of_essSurj_mapComposableArrows_two exact ⟨Octahedron.ofIso (e₁ := (e.app 0).symm) (e₂ := (e.app 1).symm) (e₃ := (e.app 2).symm) (comm₁₂ := ComposableArrows.naturality' e.inv 0 1) (comm₂₃ := ComposableArrows.naturality' e.inv 1 2) - (H := (someOctahedron rfl h₁₂' h₂₃' h₁₃').map F) _ _ _ _ _⟩ + (H := (someOctahedron rfl h₁₂' h₂₃' h₁₃').map F) ..⟩ end CategoryTheory diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index 7c5f0dc5ba024..213220d6ea296 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -43,7 +43,7 @@ the size of the biggest 3AP-free subset of `{0, ..., n - 1}`. 3AP-free, Salem-Spencer, Roth, arithmetic progression, average, three-free -/ -open Finset Function Nat +open Finset Function open scoped Pointwise variable {F α β 𝕜 E : Type*} @@ -273,7 +273,7 @@ variable {s t} {n : ℕ} @[to_additive] theorem ThreeGPFree.le_mulRothNumber (hs : ThreeGPFree (s : Set α)) (h : s ⊆ t) : s.card ≤ mulRothNumber t := - le_findGreatest (card_le_card h) ⟨s, h, rfl, hs⟩ + Nat.le_findGreatest (card_le_card h) ⟨s, h, rfl, hs⟩ @[to_additive] theorem ThreeGPFree.mulRothNumber_eq (hs : ThreeGPFree (s : Set α)) : diff --git a/Mathlib/Combinatorics/Configuration.lean b/Mathlib/Combinatorics/Configuration.lean index aabfcf9052b58..a703f011ed4da 100644 --- a/Mathlib/Combinatorics/Configuration.lean +++ b/Mathlib/Combinatorics/Configuration.lean @@ -128,7 +128,7 @@ theorem Nondegenerate.exists_injective_of_card_le [Nondegenerate P L] [Fintype P by_cases hs₁ : s.card = 1 -- If `s = {l}`, then pick a point `p ∉ l` · obtain ⟨l, rfl⟩ := Finset.card_eq_one.mp hs₁ - obtain ⟨p, hl⟩ := exists_point l + 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 diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index ab393d095b11c..220e3d4c4f141 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -291,7 +291,7 @@ lemma count_take_firstReturn_add_one : lemma count_D_lt_count_U_of_lt_firstReturn {i : ℕ} (hi : i < p.firstReturn) : (p.toList.take (i + 1)).count D < (p.toList.take (i + 1)).count U := by have ne := not_of_lt_findIdx hi - rw [decide_eq_true_eq, ← ne_eq, getElem_range] at ne + rw [decide_eq_false_iff_not, ← ne_eq, getElem_range] at ne exact lt_of_le_of_ne (p.count_D_le_count_U (i + 1)) ne.symm @[simp] @@ -307,7 +307,7 @@ lemma firstReturn_add : (p + q).firstReturn = if p = 0 then q.firstReturn else p · intro j hj rw [take_append_eq_append_take, show j + 1 - p.toList.length = 0 by omega, take_zero, append_nil] - exact (count_D_lt_count_U_of_lt_firstReturn hj).ne' + simpa using (count_D_lt_count_U_of_lt_firstReturn hj).ne' · rw [length_range, u, length_append] exact Nat.lt_add_right _ (firstReturn_lt_length h) @@ -323,6 +323,7 @@ lemma firstReturn_nest : p.nest.firstReturn = p.toList.length + 1 := by beq_iff_eq, reduceCtorEq, ite_false, take_append_eq_append_take, show j - p.toList.length = 0 by omega, take_zero, append_nil] have := p.count_D_le_count_U j + simp only [add_zero, decide_eq_false_iff_not, ne_eq] omega · simp_rw [length_range, u, length_append, length_cons] exact Nat.lt_add_one _ diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 8b5ce0135fec7..29713f6399feb 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -52,6 +52,10 @@ namespace Finset namespace Colex variable {α : Type*} [LinearOrder α] {𝒜 𝒜₁ 𝒜₂ : Finset (Finset α)} {s t : Finset α} {r : ℕ} +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- This is important for iterating Kruskal-Katona: the shadow of an initial segment is also an initial segment. -/ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : @@ -67,7 +71,7 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : · simpa [ha] using erase_le_erase_min' hts hst.ge (mem_insert_self _ _) -- Now show that if t ≤ s - min s, there is j such that t ∪ j ≤ s -- We choose j as the smallest thing not in t - simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, and_assoc] + simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, _root_.and_assoc] simp only [toColex_inj, ofColex_toColex, ne_eq, and_imp] rintro cards' (rfl | ⟨k, hks, hkt, z⟩) -- If t = s - min s, then use j = min s so t ∪ j = s @@ -125,13 +129,17 @@ variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} namespace UV +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Applying the compression makes the set smaller in colex. This is intuitive since a portion of the set is being "shifted down" as `max U < max V`. -/ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' U hU < max' V hV) (hA : compress U V s ≠ s) : toColex (compress U V s) < toColex s := by rw [compress, ite_ne_right_iff] at hA rw [compress, if_pos hA.1, lt_iff_exists_filter_lt] - simp_rw [mem_sdiff (s := s), filter_inj, and_assoc] + simp_rw [mem_sdiff (s := s), filter_inj, _root_.and_assoc] refine ⟨_, hA.1.2 <| max'_mem _ hV, not_mem_sdiff_of_mem_right <| max'_mem _ _, fun a ha ↦ ?_⟩ have : a ∉ V := fun H ↦ ha.not_le (le_max' _ _ H) have : a ∉ U := fun H ↦ ha.not_lt ((le_max' _ _ H).trans_lt h) diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index 0cb730278bd98..30662b99dcef9 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -36,7 +36,7 @@ lemma IsHamiltonian.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective f) ( /-- A hamiltonian path visits every vertex. -/ @[simp] lemma IsHamiltonian.mem_support (hp : p.IsHamiltonian) (c : α) : c ∈ p.support := by - simp only [← List.count_pos_iff_mem, hp _, Nat.zero_lt_one] + simp only [← List.count_pos_iff, hp _, Nat.zero_lt_one] /-- Hamiltonian paths are paths. -/ lemma IsHamiltonian.isPath (hp : p.IsHamiltonian) : p.IsPath := @@ -45,7 +45,7 @@ lemma IsHamiltonian.isPath (hp : p.IsHamiltonian) : p.IsPath := /-- A path whose support contains every vertex is hamiltonian. -/ lemma IsPath.isHamiltonian_of_mem (hp : p.IsPath) (hp' : ∀ w, w ∈ p.support) : p.IsHamiltonian := fun _ ↦ - le_antisymm (List.nodup_iff_count_le_one.1 hp.support_nodup _) (List.count_pos_iff_mem.2 (hp' _)) + le_antisymm (List.nodup_iff_count_le_one.1 hp.support_nodup _) (List.count_pos_iff.2 (hp' _)) lemma IsPath.isHamiltonian_iff (hp : p.IsPath) : p.IsHamiltonian ↔ ∀ w, w ∈ p.support := ⟨(·.mem_support), hp.isHamiltonian_of_mem⟩ @@ -78,7 +78,7 @@ lemma IsHamiltonianCycle.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective (hp : p.IsHamiltonianCycle) : (p.map f).IsHamiltonianCycle where toIsCycle := hp.isCycle.map hf.injective isHamiltonian_tail := by - simp only [IsHamiltonian, support_tail, support_map, ne_eq, List.map_eq_nil, support_ne_nil, + simp only [IsHamiltonian, support_tail, support_map, ne_eq, List.map_eq_nil_iff, support_ne_nil, not_false_eq_true, List.count_tail, List.head_map, beq_iff_eq, hf.surjective.forall, hf.injective, List.count_map_of_injective] intro x diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 9a5f20e639064..f40c73152000a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -590,7 +590,7 @@ end Path namespace Walk variable {G} {p} {u v : V} {H : SimpleGraph V} -variable (p : G.Walk u v) +variable {p : G.Walk u v} protected theorem IsPath.transfer (hp) (pp : p.IsPath) : (p.transfer H hp).IsPath := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index d0c4c37b647fa..e74fad0ddff83 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -32,6 +32,10 @@ namespace Finpartition variable {α : Type*} [DecidableEq α] {s t : Finset α} {m n a b : ℕ} {P : Finpartition s} +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, we can find a new partition `Q` of `s` where each part has size `m` or `m + 1`, every part of `P` is the union of parts of `Q` plus at most `m` extra elements, there are `b` parts of size `m + 1` and @@ -45,8 +49,8 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : -- 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⟩ - simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, and_assoc, - sdiff_eq_empty_iff_subset, subset_iff] + simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, + _root_.and_assoc, sdiff_eq_empty_iff_subset, subset_iff] exact fun x hx a ha => ⟨{a}, mem_map_of_mem _ (P.le hx ha), singleton_subset_iff.2 ha, mem_singleton_self _⟩ -- Prove the case `m > 0` by strong induction on `s` diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index b05bb8804a63b..6c7459ec02c9a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -83,6 +83,10 @@ lemma LocallyLinear.map (f : α ↪ β) (hG : G.LocallyLinear) : (G.map f).Local · rw [← Equiv.coe_toEmbedding, ← map_symm] exact LocallyLinear.map _ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : G.EdgeDisjointTriangles ↔ ∀ ⦃e : Sym2 α⦄, ¬ e.IsDiag → {s ∈ G.cliqueSet 3 | e ∈ (s : Finset α).sym2}.Subsingleton := by @@ -92,7 +96,7 @@ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : ext s simp only [mem_sym2_iff, Sym2.mem_iff, forall_eq_or_imp, forall_eq, Set.sep_and, Set.mem_inter_iff, Set.mem_sep_iff, mem_cliqueSet_iff, Set.mem_setOf_eq, - and_and_and_comm (b := _ ∈ _), and_self, is3Clique_iff] + and_and_and_comm (b := _ ∈ _), _root_.and_self, is3Clique_iff] constructor · rintro ⟨⟨c, d, e, hcd, hce, hde, rfl⟩, hab⟩ simp only [mem_insert, mem_singleton] at hab diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index ccdc8da2c18ea..e2c574948f9c0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ import Mathlib.Combinatorics.SimpleGraph.Maps -import Mathlib.Data.List.Lemmas /-! @@ -509,7 +508,7 @@ theorem getLast_support {G : SimpleGraph V} {a b : V} (p : G.Walk a b) : theorem tail_support_append {u v w : V} (p : G.Walk u v) (p' : G.Walk v w) : (p.append p').support.tail = p.support.tail ++ p'.support.tail := by - rw [support_append, List.tail_append_of_ne_nil _ _ (support_ne_nil _)] + rw [support_append, List.tail_append_of_ne_nil (support_ne_nil _)] theorem support_eq_cons {u v : V} (p : G.Walk u v) : p.support = u :: p.support.tail := by cases p <;> simp diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 79b3e2892c891..4e120e2d50249 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -436,7 +436,7 @@ theorem rowLens_length_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (hpo (ofRowLens w hw).rowLens.length = w.length := by simp only [length_rowLens, colLen, Nat.find_eq_iff, mem_cells, mem_ofRowLens, lt_self_iff_false, IsEmpty.exists_iff, Classical.not_not] - exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.getElem_mem _ _ hn)⟩⟩ + exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.getElem_mem hn)⟩⟩ /-- The length of the `i`th row in `ofRowLens w hw` is the `i`th entry of `w` -/ theorem rowLen_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (i : Fin w.length) : diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 1b60a716c05f3..7017d99c1eb68 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -75,20 +75,20 @@ theorem ack_succ_succ (m n : ℕ) : ack (m + 1) (n + 1) = ack m (ack (m + 1) n) @[simp] theorem ack_one (n : ℕ) : ack 1 n = n + 2 := by induction' n with n IH - · rfl + · simp · simp [IH] @[simp] theorem ack_two (n : ℕ) : ack 2 n = 2 * n + 3 := by induction' n with n IH - · rfl + · simp · simpa [mul_succ] -- Porting note: re-written to get rid of ack_three_aux @[simp] theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by induction' n with n IH - · rfl + · simp · rw [ack_succ_succ, IH, ack_two, Nat.succ_add, Nat.pow_succ 2 (n + 3), mul_comm _ 2, Nat.mul_sub_left_distrib, ← Nat.sub_add_comm, two_mul 3, Nat.add_sub_add_right] have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 084c5246a26c6..a2de5176d359b 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -223,7 +223,7 @@ theorem rice₂ (C : Set Code) (H : ∀ cf cg, eval cf = eval cg → (cf ∈ C (Partrec.nat_iff.1 <| eval_part.comp (const cg) Computable.id) ((hC _).1 fC), fun h => by { obtain rfl | rfl := h <;> simpa [ComputablePred, Set.mem_empty_iff_false] using - ⟨⟨inferInstance⟩, Computable.const _⟩ }⟩ + Computable.const _}⟩ /-- The Halting problem is recursively enumerable -/ theorem halting_problem_re (n) : RePred fun c => (eval c n).Dom := @@ -281,8 +281,6 @@ namespace Nat.Partrec' open Mathlib.Vector Partrec Computable -open Nat (Partrec') - open Nat.Partrec' theorem to_part {n f} (pf : @Partrec' n f) : _root_.Partrec f := by diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index cc112e2e09266..a8779cd1c9283 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -159,7 +159,7 @@ lemma mem_kstar_iff_exists_nonempty {x : List α} : x ∈ l∗ ↔ ∃ S : List (List α), x = S.join ∧ ∀ y ∈ S, y ∈ l ∧ y ≠ [] := by constructor · rintro ⟨S, rfl, h⟩ - refine ⟨S.filter fun l ↦ !List.isEmpty l, by simp, fun y hy ↦ ?_⟩ + refine ⟨S.filter fun l ↦ !List.isEmpty l, by simp [List.join_filter_not_isEmpty], fun y hy ↦ ?_⟩ -- Porting note: The previous code was: -- rw [mem_filter, empty_iff_eq_nil] at hy rw [mem_filter, Bool.not_eq_true', ← Bool.bool_iff_false, List.isEmpty_iff] at hy diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index f563f6babf2c4..74b64e82859ab 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -170,7 +170,7 @@ private theorem encode_ofNatCode : ∀ n, encodeCode (ofNatCode n) = n instance instDenumerable : Denumerable Code := mk' ⟨encodeCode, ofNatCode, fun c => by - induction c <;> try {rfl} <;> simp [encodeCode, ofNatCode, Nat.div2_val, *], + induction c <;> simp [encodeCode, ofNatCode, Nat.div2_val, *], encode_ofNatCode⟩ theorem encodeCode_eq : encode = encodeCode := @@ -900,7 +900,7 @@ private theorem hG : Primrec G := by Primrec.fst private theorem evaln_map (k c n) : - ((((List.range k)[n]?).map (evaln k c)).bind fun b => b) = evaln k c n := by + ((List.range k)[n]?.bind fun a ↦ evaln k c a) = evaln k c n := by by_cases kn : n < k · simp [List.getElem?_range kn] · rw [List.getElem?_len_le] @@ -937,7 +937,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a (List.range n.unpair.1).map (evaln n.unpair.1 (ofNat Code n.unpair.2))) (k', c') n = evaln k' c' n := by intro k₁ c₁ n₁ hl - simp [lup, List.getElem?_range hl, evaln_map, Bind.bind] + simp [lup, List.getElem?_range hl, evaln_map, Bind.bind, Option.bind_map] cases' c with cf cg cf cg cf cg cf <;> simp [evaln, nk, Bind.bind, Functor.map, Seq.seq, pure] · cases' encode_lt_pair cf cg with lf lg @@ -969,7 +969,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a (Primrec.option_bind (Primrec.list_get?.comp (this.comp (_root_.Primrec.const ()) (Primrec.encode_iff.2 Primrec.fst)) Primrec.snd) Primrec.snd.to₂).of_eq - fun ⟨⟨k, c⟩, n⟩ => by simp [evaln_map] + fun ⟨⟨k, c⟩, n⟩ => by simp [evaln_map, Option.bind_map] end diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index ea36cdcbb6053..73b365b5200ee 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -263,7 +263,7 @@ def ListBlank.nth {Γ} [Inhabited Γ] (l : ListBlank Γ) (n : ℕ) : Γ := by rw [List.getI_eq_default _ h] rcases le_or_lt _ n with h₂ | h₂ · rw [List.getI_eq_default _ h₂] - rw [List.getI_eq_get _ h₂, List.get_eq_getElem, List.getElem_append_right' h, + rw [List.getI_eq_get _ h₂, List.get_eq_getElem, List.getElem_append_right h, List.getElem_replicate] @[simp] diff --git a/Mathlib/Control/Applicative.lean b/Mathlib/Control/Applicative.lean index 87acbfcaf5833..3b6bb33ac8951 100644 --- a/Mathlib/Control/Applicative.lean +++ b/Mathlib/Control/Applicative.lean @@ -5,6 +5,7 @@ Authors: Simon Hudon -/ import Mathlib.Algebra.Group.Defs import Mathlib.Control.Functor +import Mathlib.Control.Basic /-! # `applicative` instances @@ -28,7 +29,7 @@ variable {α β γ σ : Type u} theorem Applicative.map_seq_map (f : α → β → γ) (g : σ → β) (x : F α) (y : F σ) : f <$> x <*> g <$> y = ((· ∘ g) ∘ f) <$> x <*> y := by - simp [flip, functor_norm] + simp [flip, functor_norm, Function.comp_def] theorem Applicative.pure_seq_eq_map' (f : α → β) : ((pure f : F (α → β)) <*> ·) = (f <$> ·) := by ext; simp [functor_norm] diff --git a/Mathlib/Control/Basic.lean b/Mathlib/Control/Basic.lean index 9f03e34b31c1e..e6e6aa94abeee 100644 --- a/Mathlib/Control/Basic.lean +++ b/Mathlib/Control/Basic.lean @@ -18,12 +18,7 @@ variable {α β γ : Type u} section Functor -variable {f : Type u → Type v} [Functor f] [LawfulFunctor f] -@[functor_norm] -theorem Functor.map_map (m : α → β) (g : β → γ) (x : f α) : g <$> m <$> x = (g ∘ m) <$> x := - (comp_map _ _ _).symm --- order of implicits --- order of implicits +attribute [functor_norm] Functor.map_map end Functor @@ -67,10 +62,6 @@ section Monad variable {m : Type u → Type v} [Monad m] [LawfulMonad m] -theorem map_bind (x : m α) {g : α → m β} {f : β → γ} : - f <$> (x >>= g) = x >>= fun a => f <$> g a := by - rw [← bind_pure_comp, bind_assoc]; simp [bind_pure_comp] - theorem seq_bind_eq (x : m α) {g : β → m γ} {f : α → β} : f <$> x >>= g = x >>= g ∘ f := show bind (f <$> x) g = bind x (g ∘ f) by diff --git a/Mathlib/Control/Functor.lean b/Mathlib/Control/Functor.lean index 59ebe2553d456..305a88d2cc83f 100644 --- a/Mathlib/Control/Functor.lean +++ b/Mathlib/Control/Functor.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ -import Mathlib.Control.Basic +import Mathlib.Tactic.Attr.Register import Mathlib.Data.Set.Defs import Mathlib.Tactic.TypeStar import Batteries.Tactic.Lint @@ -181,7 +181,7 @@ protected theorem id_map : ∀ x : Comp F G α, Comp.map id x = x protected theorem comp_map (g' : α → β) (h : β → γ) : ∀ x : Comp F G α, Comp.map (h ∘ g') x = Comp.map h (Comp.map g' x) - | Comp.mk x => by simp [Comp.map, Comp.mk, Functor.map_comp_map, functor_norm] + | Comp.mk x => by simp [Comp.map, Comp.mk, Functor.map_comp_map, functor_norm, Function.comp_def] -- Porting note: `Comp.mk` wasn't needed in mathlib3 instance lawfulFunctor : LawfulFunctor (Comp F G) where diff --git a/Mathlib/Control/Traversable/Basic.lean b/Mathlib/Control/Traversable/Basic.lean index e9103c5c55aa4..52d25c007718a 100644 --- a/Mathlib/Control/Traversable/Basic.lean +++ b/Mathlib/Control/Traversable/Basic.lean @@ -6,6 +6,7 @@ Authors: Simon Hudon import Mathlib.Data.Option.Defs import Mathlib.Control.Functor import Batteries.Data.List.Basic +import Mathlib.Control.Basic /-! # Traversable type class diff --git a/Mathlib/Control/Traversable/Equiv.lean b/Mathlib/Control/Traversable/Equiv.lean index fbd1a7fcc724e..db32b26b14bb2 100644 --- a/Mathlib/Control/Traversable/Equiv.lean +++ b/Mathlib/Control/Traversable/Equiv.lean @@ -53,7 +53,7 @@ protected theorem id_map {α : Type u} (x : t' α) : Equiv.map eqv id x = x := b protected theorem comp_map {α β γ : Type u} (g : α → β) (h : β → γ) (x : t' α) : Equiv.map eqv (h ∘ g) x = Equiv.map eqv h (Equiv.map eqv g x) := by - simpa [Equiv.map] using comp_map .. + simp [Equiv.map, Function.comp_def] protected theorem lawfulFunctor : @LawfulFunctor _ (Equiv.functor eqv) := -- Porting note: why is `_inst` required here? diff --git a/Mathlib/Control/Traversable/Instances.lean b/Mathlib/Control/Traversable/Instances.lean index 9e4afad602947..a91d6a6202f43 100644 --- a/Mathlib/Control/Traversable/Instances.lean +++ b/Mathlib/Control/Traversable/Instances.lean @@ -32,7 +32,7 @@ theorem Option.id_traverse {α} (x : Option α) : Option.traverse (pure : α → theorem Option.comp_traverse {α β γ} (f : β → F γ) (g : α → G β) (x : Option α) : Option.traverse (Comp.mk ∘ (f <$> ·) ∘ g) x = Comp.mk (Option.traverse f <$> Option.traverse g x) := by - cases x <;> simp! [functor_norm] <;> rfl + cases x <;> (simp! [functor_norm] <;> rfl) theorem Option.traverse_eq_map_id {α β} (f : α → β) (x : Option α) : Option.traverse ((pure : _ → Id _) ∘ f) x = (pure : _ → Id _) (f <$> x) := by cases x <;> rfl @@ -148,7 +148,7 @@ variable [LawfulApplicative G] protected theorem comp_traverse {α β γ : Type u} (f : β → F γ) (g : α → G β) (x : σ ⊕ α) : Sum.traverse (Comp.mk ∘ (f <$> ·) ∘ g) x = Comp.mk.{u} (Sum.traverse f <$> Sum.traverse g x) := by - cases x <;> simp! [Sum.traverse, map_id, functor_norm] <;> rfl + cases x <;> (simp! [Sum.traverse, map_id, functor_norm] <;> rfl) protected theorem traverse_eq_map_id {α β} (f : α → β) (x : σ ⊕ α) : Sum.traverse ((pure : _ → Id _) ∘ f) x = (pure : _ → Id _) (f <$> x) := by diff --git a/Mathlib/Data/Array/ExtractLemmas.lean b/Mathlib/Data/Array/ExtractLemmas.lean index bc66fc0660fbc..b27b5245af41d 100644 --- a/Mathlib/Data/Array/ExtractLemmas.lean +++ b/Mathlib/Data/Array/ExtractLemmas.lean @@ -27,7 +27,7 @@ theorem extract_append_left {a b : Array α} {i j : Nat} (h : j ≤ a.size) : · simp only [size_extract, size_append] omega · intro h1 h2 h3 - rw [get_extract, get_append_left, get_extract] + rw [getElem_extract, getElem_append_left, getElem_extract] theorem extract_append_right {a b : Array α} {i j : Nat} (h : a.size ≤ i) : (a ++ b).extract i j = b.extract (i - a.size) (j - a.size) := by @@ -35,8 +35,8 @@ theorem extract_append_right {a b : Array α} {i j : Nat} (h : a.size ≤ i) : · rw [size_extract, size_extract, size_append] omega · intro k hi h2 - rw [get_extract, get_extract, - get_append_right (show size a ≤ i + k by omega)] + rw [getElem_extract, getElem_extract, + getElem_append_right (show size a ≤ i + k by omega)] congr omega @@ -50,6 +50,6 @@ theorem extract_extract {s1 e2 e1 s2 : Nat} {a : Array α} (h : s1 + e2 ≤ e1) · simp only [size_extract] omega · intro i h1 h2 - simp only [get_extract, Nat.add_assoc] + simp only [getElem_extract, Nat.add_assoc] end Array diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index db1c8bc2bdd5f..ebb4970f7a3ce 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -44,7 +44,14 @@ lemma le_def : f ≤ g ↔ ∀ i, f i ≤ g i := Iff.rfl def orderEmbeddingToFun : (Π₀ i, α i) ↪o ∀ i, α i where toFun := DFunLike.coe inj' := DFunLike.coe_injective - map_rel_iff' := by rfl + map_rel_iff' := + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + Iff.rfl @[simp, norm_cast] lemma coe_orderEmbeddingToFun : ⇑(orderEmbeddingToFun (α := α)) = DFunLike.coe := rfl diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index d4c6eeb432941..d875809cf1a4a 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -414,16 +414,6 @@ section Monoid 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)] --- Porting note (#10618): removing `simp`, `simp` can prove it with AddCommMonoid instance -protected theorem zero_add [NeZero n] (k : Fin n) : 0 + k = k := by - simp [Fin.ext_iff, add_def, mod_eq_of_lt (is_lt k)] - -instance {a : ℕ} [NeZero n] : OfNat (Fin n) a where - ofNat := Fin.ofNat' a n.pos_of_neZero - -instance inhabited (n : ℕ) [NeZero n] : Inhabited (Fin n) := - ⟨0⟩ - instance inhabitedFinOneAdd (n : ℕ) : Inhabited (Fin (1 + n)) := haveI : NeZero (1 + n) := by rw [Nat.add_comm]; infer_instance inferInstance @@ -434,8 +424,8 @@ theorem default_eq_zero (n : ℕ) [NeZero n] : (default : Fin n) = 0 := section from_ad_hoc -@[simp] lemma ofNat'_zero {h : 0 < n} [NeZero n] : (Fin.ofNat' 0 h : Fin n) = 0 := rfl -@[simp] lemma ofNat'_one {h : 0 < n} [NeZero n] : (Fin.ofNat' 1 h : Fin n) = 1 := rfl +@[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 @@ -516,13 +506,6 @@ lemma natCast_strictMono (hbn : b ≤ n) (hab : a < b) : (a : Fin (n + 1)) < b : end OfNatCoe -@[simp] -theorem one_eq_zero_iff [NeZero n] : (1 : Fin n) = 0 ↔ n = 1 := by - obtain _ | _ | n := n <;> simp [Fin.ext_iff] - -@[simp] -theorem zero_eq_one_iff [NeZero n] : (0 : Fin n) = 1 ↔ n = 1 := by rw [eq_comm, one_eq_zero_iff] - end Add section Succ @@ -578,10 +561,6 @@ This one instead uses a `NeZero n` typeclass hypothesis. theorem le_zero_iff' {n : ℕ} [NeZero n] {k : Fin n} : k ≤ 0 ↔ k = 0 := ⟨fun h => Fin.ext <| by rw [Nat.eq_zero_of_le_zero h]; rfl, by rintro rfl; exact Nat.le_refl _⟩ --- Move to Batteries? -@[simp] theorem cast_refl {n : Nat} (h : n = n) : - Fin.cast h = id := rfl - -- TODO: Move to Batteries @[simp] lemma castLE_inj {hmn : m ≤ n} {a b : Fin m} : castLE hmn a = castLE hmn b ↔ a = b := by simp [Fin.ext_iff] @@ -775,7 +754,7 @@ theorem castSucc_ne_zero_of_lt {p i : Fin n} (h : p < i) : castSucc i ≠ 0 := b exact ((zero_le _).trans_lt h).ne' theorem succ_ne_last_iff (a : Fin (n + 1)) : succ a ≠ last (n + 1) ↔ a ≠ last n := - not_iff_not.mpr <| succ_eq_last_succ a + not_iff_not.mpr <| succ_eq_last_succ theorem succ_ne_last_of_lt {p i : Fin n} (h : i < p) : succ i ≠ last n := by cases n @@ -838,7 +817,7 @@ theorem le_pred_iff {j : Fin n} {i : Fin (n + 1)} (hi : i ≠ 0) : j ≤ pred i rw [← succ_le_succ_iff, succ_pred] theorem castSucc_pred_eq_pred_castSucc {a : Fin (n + 1)} (ha : a ≠ 0) - (ha' := a.castSucc_ne_zero_iff.mpr ha) : + (ha' := castSucc_ne_zero_iff.mpr ha) : (a.pred ha).castSucc = (castSucc a).pred ha' := rfl theorem castSucc_pred_add_one_eq {a : Fin (n + 1)} (ha : a ≠ 0) : @@ -1491,7 +1470,7 @@ theorem eq_zero (n : Fin 1) : n = 0 := Subsingleton.elim _ _ instance uniqueFinOne : Unique (Fin 1) where uniq _ := Subsingleton.elim _ _ -@[simp] +@[deprecated val_eq_zero (since := "2024-09-18")] theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a 0] lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := by diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 4d2f060f732a2..e4cb2370272a7 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -788,7 +788,7 @@ theorem insertNth_apply_succAbove (i : Fin (n + 1)) (x : α i) (p : ∀ j, α (i generalize hk : castPred ((succAbove i) j) H₁ = k rw [castPred_succAbove _ _ hlt] at hk; cases hk intro; rfl - · generalize_proofs H₁ H₂; revert H₂ + · generalize_proofs H₀ H₁ H₂; revert H₂ generalize hk : pred (succAbove i j) H₁ = k erw [pred_succAbove _ _ (Fin.not_lt.1 hlt)] at hk; cases hk intro; rfl diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 7f72be2b670cb..7f311baa1744c 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -123,7 +123,7 @@ theorem antidiagonalTuple_one (n : ℕ) : antidiagonalTuple 1 n = [![n]] := by Nat.sub_self, List.bind_append, List.bind_singleton, List.bind_map] conv_rhs => rw [← List.nil_append [![n]]] congr 1 - simp_rw [List.bind_eq_nil, List.mem_range, List.map_eq_nil] + simp_rw [List.bind_eq_nil_iff, List.mem_range, List.map_eq_nil_iff] intro x hx obtain ⟨m, rfl⟩ := Nat.exists_eq_add_of_lt hx rw [add_assoc, add_tsub_cancel_left, antidiagonalTuple_zero_succ] diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 86d298bb40842..b84e46a5f7a05 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -125,7 +125,7 @@ assert_not_exists CompleteLattice assert_not_exists OrderedCommMonoid -open Multiset Subtype Nat Function +open Multiset Subtype Function universe u @@ -2518,6 +2518,8 @@ end Filter section Range +open Nat + variable {n m l : ℕ} /-- `range n` is the set of natural numbers less than `n`. -/ diff --git a/Mathlib/Data/Finset/Functor.lean b/Mathlib/Data/Finset/Functor.lean index 4319e4a6b1b3a..1359446c44d7e 100644 --- a/Mathlib/Data/Finset/Functor.lean +++ b/Mathlib/Data/Finset/Functor.lean @@ -198,11 +198,16 @@ theorem map_comp_coe (h : α → β) : Functor.map h ∘ Multiset.toFinset = Multiset.toFinset ∘ Functor.map h := funext fun _ => image_toFinset +@[simp] +theorem map_comp_coe_apply (h : α → β) (s : Multiset α) : + s.toFinset.image h = (h <$> s).toFinset := + congrFun (map_comp_coe h) s + theorem map_traverse (g : α → G β) (h : β → γ) (s : Finset α) : Functor.map h <$> traverse g s = traverse (Functor.map h ∘ g) s := by unfold traverse - simp only [map_comp_coe, functor_norm] - rw [LawfulFunctor.comp_map, Multiset.map_traverse] + simp only [Functor.map_map, fmap_def, map_comp_coe_apply, Multiset.fmap_def, ← + Multiset.map_traverse] end Traversable diff --git a/Mathlib/Data/Finset/NatDivisors.lean b/Mathlib/Data/Finset/NatDivisors.lean index cbc7aa896c2f5..03edf844dbd26 100644 --- a/Mathlib/Data/Finset/NatDivisors.lean +++ b/Mathlib/Data/Finset/NatDivisors.lean @@ -16,12 +16,16 @@ exhibiting `Nat.divisors` as a multiplicative homomorphism from `ℕ` to `Finset open Nat Finset open scoped Pointwise +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The divisors of a product of natural numbers are the pointwise product of the divisors of the factors. -/ lemma Nat.divisors_mul (m n : ℕ) : divisors (m * n) = divisors m * divisors n := by ext k simp_rw [mem_mul, mem_divisors, dvd_mul, mul_ne_zero_iff, ← exists_and_left, ← exists_and_right] - simp only [and_assoc, and_comm, and_left_comm] + simp only [_root_.and_assoc, _root_.and_comm, and_left_comm] /-- `Nat.divisors` as a `MonoidHom`. -/ @[simps] diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index fd8ebbddcb92b..45c0a90562ef5 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -48,8 +48,6 @@ assert_not_exists MulAction open Function -open Nat - universe u v variable {α β γ : Type*} diff --git a/Mathlib/Data/Fintype/Prod.lean b/Mathlib/Data/Fintype/Prod.lean index 0d91984495e7f..2a5dfcdbbbf59 100644 --- a/Mathlib/Data/Fintype/Prod.lean +++ b/Mathlib/Data/Fintype/Prod.lean @@ -14,8 +14,6 @@ import Mathlib.Data.Finset.Prod open Function -open Nat - universe u v variable {α β γ : Type*} diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index b5619f8bcea08..4cb4e2f0e0ba6 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -43,7 +43,13 @@ protected lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := ⟨fun h ↦ ⟨Int.le_of_eq h, Int.ge_of_eq h⟩, fun h ↦ Int.le_antisymm h.1 h.2⟩ protected lemma le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := by rw [Int.le_antisymm_iff, Int.lt_iff_le_not_le, ← and_or_left]; simp [em] -protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, or_comm] + +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ +protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, + _root_.or_comm] end Order diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index b4562695ae718..87990322c4065 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -3,6 +3,7 @@ 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.Control.Basic import Mathlib.Data.Nat.Defs import Mathlib.Data.Option.Basic import Mathlib.Data.List.Defs @@ -35,8 +36,6 @@ variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : Lis @[deprecated (since := "2024-07-27")] theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl -@[deprecated (since := "2024-06-07")] alias toArray_data := Array.data_toArray - -- Porting note: Delete this attribute -- attribute [inline] List.head! @@ -59,9 +58,6 @@ instance : Std.Associative (α := List α) Append.append where theorem singleton_injective : Injective fun a : α => [a] := fun _ _ h => (cons_eq_cons.1 h).1 -theorem singleton_inj {a b : α} : [a] = [b] ↔ a = b := - singleton_injective.eq_iff - theorem set_of_mem_cons (l : List α) (a : α) : { x | x ∈ a :: l } = insert a { x | x ∈ l } := Set.ext fun _ => mem_cons @@ -199,10 +195,6 @@ theorem map_subset_iff {l₁ l₂ : List α} (f : α → β) (h : Injective f) : theorem append_eq_has_append {L₁ L₂ : List α} : List.append L₁ L₂ = L₁ ++ L₂ := rfl -@[deprecated (since := "2024-03-24")] alias append_eq_cons_iff := append_eq_cons - -@[deprecated (since := "2024-03-24")] alias cons_eq_append_iff := cons_eq_append - @[deprecated (since := "2024-01-18")] alias append_left_cancel := append_cancel_left @[deprecated (since := "2024-01-18")] alias append_right_cancel := append_cancel_right @@ -229,10 +221,10 @@ theorem replicate_subset_singleton (n) (a : α) : replicate n a ⊆ [a] := fun _ mem_singleton.2 (eq_of_mem_replicate h) theorem subset_singleton_iff {a : α} {L : List α} : L ⊆ [a] ↔ ∃ n, L = replicate n a := by - simp only [eq_replicate, subset_def, mem_singleton, exists_eq_left'] + simp only [eq_replicate_iff, subset_def, mem_singleton, exists_eq_left'] theorem replicate_right_injective {n : ℕ} (hn : n ≠ 0) : Injective (@replicate α n) := - fun _ _ h => (eq_replicate.1 h).2 _ <| mem_replicate.2 ⟨hn, rfl⟩ + fun _ _ h => (eq_replicate_iff.1 h).2 _ <| mem_replicate.2 ⟨hn, rfl⟩ theorem replicate_right_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : replicate n a = replicate n b ↔ a = b := @@ -360,21 +352,10 @@ lemma getLast_filter {p : α → Bool} : /-! ### getLast? -/ --- This is a duplicate of `getLast?_eq_none_iff`. --- We should remove one of them. -theorem getLast?_eq_none : ∀ {l : List α}, getLast? l = none ↔ l = [] - | [] => by simp - | [a] => by simp - | a :: b :: l => by simp [@getLast?_eq_none (b :: l)] +@[deprecated (since := "2024-09-06")] alias getLast?_eq_none := getLast?_eq_none_iff @[deprecated (since := "2024-06-20")] alias getLast?_isNone := getLast?_eq_none -@[simp] -theorem getLast?_isSome : ∀ {l : List α}, l.getLast?.isSome ↔ l ≠ [] - | [] => by simp - | [a] => by simp - | a :: b :: l => by simp [@getLast?_isSome (b :: l)] - theorem mem_getLast?_eq_getLast : ∀ {l : List α} {x : α}, x ∈ l.getLast? → ∃ h, x = getLast l h | [], x, hx => False.elim <| by simp at hx | [a], x, hx => @@ -395,10 +376,6 @@ theorem mem_getLast?_cons {x y : α} : ∀ {l : List α}, x ∈ l.getLast? → x | [], _ => by contradiction | _ :: _, h => h -theorem mem_of_mem_getLast? {l : List α} {a : α} (ha : a ∈ l.getLast?) : a ∈ l := - let ⟨_, h₂⟩ := mem_getLast?_eq_getLast ha - h₂.symm ▸ getLast_mem _ - theorem dropLast_append_getLast? : ∀ {l : List α}, ∀ a ∈ l.getLast?, dropLast l ++ [a] = l | [], a, ha => (Option.not_mem_none a ha).elim | [a], _, rfl => rfl @@ -462,9 +439,6 @@ theorem eq_cons_of_mem_head? {x : α} : ∀ {l : List α}, x ∈ l.head? → l = simp only [head?, Option.mem_def, Option.some_inj] at h exact h ▸ rfl -theorem mem_of_mem_head? {x : α} {l : List α} (h : x ∈ l.head?) : x ∈ l := - (eq_cons_of_mem_head? h).symm ▸ mem_cons_self _ _ - @[simp] theorem head!_cons [Inhabited α] (a : α) (l : List α) : head! (a :: l) = a := rfl @[simp] @@ -960,9 +934,6 @@ theorem infix_bind_of_mem {a : α} {as : List α} (h : a ∈ as) (f : α → Lis theorem map_eq_map {α β} (f : α → β) (l : List α) : f <$> l = map f l := rfl -@[simp] -theorem map_tail (f : α → β) (l) : map f (tail l) = tail (map f l) := by cases l <;> rfl - /-- A single `List.map` of a composition of functions is equal to composing a `List.map` with another `List.map`, fully applied. This is the reverse direction of `List.map_map`. @@ -1198,7 +1169,7 @@ lemma append_cons_inj_of_not_mem {x₁ x₂ z₁ z₂ : List α} {a₁ a₂ : α (notin_x : a₂ ∉ x₁) (notin_z : a₂ ∉ z₁) : x₁ ++ a₁ :: z₁ = x₂ ++ a₂ :: z₂ ↔ x₁ = x₂ ∧ a₁ = a₂ ∧ z₁ = z₂ := by constructor - · simp only [append_eq_append_iff, cons_eq_append, cons_eq_cons] + · simp only [append_eq_append_iff, cons_eq_append_iff, cons_eq_cons] rintro (⟨c, rfl, ⟨rfl, rfl, rfl⟩ | ⟨d, rfl, rfl⟩⟩ | ⟨c, rfl, ⟨rfl, rfl, rfl⟩ | ⟨d, rfl, rfl⟩⟩) <;> simp_all · rintro ⟨rfl, rfl, rfl⟩ @@ -1267,7 +1238,7 @@ theorem getElem_succ_scanl {i : ℕ} (h : i + 1 < (scanl f b l).length) : · simp only [length] at h exact absurd h (by omega) · simp_rw [scanl_cons] - rw [getElem_append_right'] + rw [getElem_append_right] · simp only [length, Nat.zero_add 1, succ_add_sub_one, hi]; rfl · simp only [length_singleton]; omega @@ -1360,25 +1331,12 @@ local notation a " ⋆ " b => op a b /-- Notation for `foldl op a l`. -/ local notation l " <*> " a => foldl op a l -theorem foldl_assoc : ∀ {l : List α} {a₁ a₂}, (l <*> a₁ ⋆ a₂) = a₁ ⋆ l <*> a₂ - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => - calc - ((a :: l) <*> a₁ ⋆ a₂) = l <*> a₁ ⋆ a₂ ⋆ a := by simp only [foldl_cons, ha.assoc] - _ = a₁ ⋆ (a :: l) <*> a₂ := by rw [foldl_assoc, foldl_cons] - theorem foldl_op_eq_op_foldr_assoc : ∀ {l : List α} {a₁ a₂}, ((l <*> a₁) ⋆ a₂) = a₁ ⋆ l.foldr (· ⋆ ·) a₂ | [], a₁, a₂ => rfl | a :: l, a₁, a₂ => by simp only [foldl_cons, foldr_cons, foldl_assoc, ha.assoc]; rw [foldl_op_eq_op_foldr_assoc] -theorem foldr_assoc : ∀ {l : List α} {a₁ a₂}, l.foldr op (op a₁ a₂) = op (l.foldr op a₁) a₂ - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => by - simp only [foldr_cons, ha.assoc] - rw [foldr_assoc] - variable [hc : Std.Commutative op] theorem foldl_assoc_comm_cons {l : List α} {a₁ a₂} : ((a₁ :: l) <*> a₂) = a₁ ⋆ l <*> a₂ := by @@ -1555,18 +1513,19 @@ theorem modifyLast.go_append_one (f : α → α) (a : α) (tl : List α) (r : Ar rw [modifyLast.go, modifyLast.go] case x_3 | x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one _ _ tl _, modifyLast.go_append_one _ _ tl (Array.push #[] hd)] - simp only [Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, append_assoc] + simp only [Array.toListAppend_eq, Array.push_toList, Array.toList_toArray, nil_append, + append_assoc] theorem modifyLast_append_one (f : α → α) (a : α) (l : List α) : modifyLast f (l ++ [a]) = l ++ [f a] := by cases l with | nil => - simp only [nil_append, modifyLast, modifyLast.go, Array.toListAppend_eq, Array.data_toArray] + simp only [nil_append, modifyLast, modifyLast.go, Array.toListAppend_eq, Array.toList_toArray] | cons _ tl => simp only [cons_append, modifyLast] rw [modifyLast.go] case x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) - rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_data, Array.data_toArray, + rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_toList, Array.toList_toArray, nil_append, cons_append, nil_append, cons_inj_right] exact modifyLast_append_one _ _ tl @@ -1617,11 +1576,13 @@ variable (f : α → Option α) theorem lookmap.go_append (l : List α) (acc : Array α) : lookmap.go f l acc = acc.toListAppend (lookmap f l) := by cases l with - | nil => rfl + | nil => simp [go, lookmap] | cons hd tl => rw [lookmap, go, go] cases f hd with - | none => simp only [go_append tl _, Array.toListAppend_eq, append_assoc, Array.push_data]; rfl + | none => + simp only [go_append tl _, Array.toListAppend_eq, append_assoc, Array.push_toList] + rfl | some a => rfl @[simp] @@ -1631,13 +1592,13 @@ theorem lookmap_nil : [].lookmap f = [] := @[simp] theorem lookmap_cons_none {a : α} (l : List α) (h : f a = none) : (a :: l).lookmap f = a :: l.lookmap f := by - simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.data_toArray, nil_append] + simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.toList_toArray, nil_append] rw [lookmap.go_append, h]; rfl @[simp] theorem lookmap_cons_some {a b : α} (l : List α) (h : f a = some b) : (a :: l).lookmap f = b :: l := by - simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.data_toArray, nil_append] + simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.toList_toArray, nil_append] rw [h] theorem lookmap_some : ∀ l : List α, l.lookmap some = l @@ -1798,7 +1759,7 @@ lemma filter_attach (l : List α) (p : α → Bool) : ← filter_map, attach_map_subtype_val] lemma filter_comm (q) (l : List α) : filter p (filter q l) = filter q (filter p l) := by - simp [and_comm] + simp [Bool.and_comm] @[simp] theorem filter_true (l : List α) : @@ -1920,7 +1881,7 @@ theorem erase_getElem [DecidableEq ι] {l : List ι} {i : ℕ} (hi : i < l.lengt | succ i => have hi' : i < l.length := by simpa using hi if ha : a = l[i] then - simpa [ha] using .trans (perm_cons_erase (l.getElem_mem i _)) (.cons _ (IH hi')) + simpa [ha] using .trans (perm_cons_erase (getElem_mem _)) (.cons _ (IH hi')) else simpa [ha] using IH hi' @@ -2240,24 +2201,21 @@ end Forall /-! ### Miscellaneous lemmas -/ -@[simp] -theorem getElem_attach (L : List α) (i : Nat) (h : i < L.attach.length) : - L.attach[i].1 = L[i]'(length_attach L ▸ h) := - calc - L.attach[i].1 = (L.attach.map Subtype.val)[i]'(by simpa using h) := by - rw [getElem_map] - _ = L[i]'_ := by congr 2; simp - theorem get_attach (L : List α) (i) : - (L.attach.get i).1 = L.get ⟨i, length_attach L ▸ i.2⟩ := by simp + (L.attach.get i).1 = L.get ⟨i, length_attach (L := L) ▸ i.2⟩ := by simp +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp 1100] theorem mem_map_swap (x : α) (y : β) (xs : List (α × β)) : (y, x) ∈ map Prod.swap xs ↔ (x, y) ∈ xs := by induction' xs with x xs xs_ih · simp only [not_mem_nil, map_nil] · cases' x with a b - simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, and_comm] + simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, + _root_.and_comm] theorem dropSlice_eq (xs : List α) (n m : ℕ) : dropSlice n m xs = xs.take n ++ xs.drop (n + m) := by induction n generalizing xs diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index a1953d0529518..1213160f03d0a 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -47,6 +47,10 @@ theorem Chain.iff_mem {a : α} {l : List α} : theorem chain_singleton {a b : α} : Chain R a [b] ↔ R a b := by simp only [chain_cons, Chain.nil, and_true] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem chain_split {a b : α} {l₁ l₂ : List α} : Chain R a (l₁ ++ b :: l₂) ↔ Chain R a (l₁ ++ [b]) ∧ Chain R b l₂ := by induction' l₁ with x l₁ IH generalizing a <;> @@ -226,14 +230,18 @@ theorem Chain'.cons' {x} : ∀ {l : List α}, Chain' R l → (∀ y ∈ l.head?, theorem chain'_cons' {x l} : Chain' R (x :: l) ↔ (∀ y ∈ head? l, R x y) ∧ Chain' R l := ⟨fun h => ⟨h.rel_head?, h.tail⟩, fun ⟨h₁, h₂⟩ => h₂.cons' h₁⟩ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ theorem chain'_append : ∀ {l₁ l₂ : List α}, Chain' R (l₁ ++ l₂) ↔ Chain' R l₁ ∧ Chain' R l₂ ∧ ∀ x ∈ l₁.getLast?, ∀ y ∈ l₂.head?, R x y | [], l => by simp - | [a], l => by simp [chain'_cons', and_comm] + | [a], l => by simp [chain'_cons', _root_.and_comm] | a :: b :: l₁, l₂ => by rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, - and_assoc] + _root_.and_assoc] simp theorem Chain'.append (h₁ : Chain' R l₁) (h₂ : Chain' R l₂) @@ -272,12 +280,16 @@ theorem Chain'.imp_head {x y} (h : ∀ {z}, R x z → R y z) {l} (hl : Chain' R Chain' R (y :: l) := hl.tail.cons' fun _ hz => h <| hl.rel_head? hz +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem chain'_reverse : ∀ {l}, Chain' R (reverse l) ↔ Chain' (flip R) l | [] => Iff.rfl | [a] => by simp only [chain'_singleton, reverse_singleton] | a :: b :: l => by rw [chain'_cons, reverse_cons, reverse_cons, append_assoc, cons_append, nil_append, - chain'_split, ← reverse_cons, @chain'_reverse (b :: l), and_comm, chain'_pair, flip] + chain'_split, ← reverse_cons, @chain'_reverse (b :: l), _root_.and_comm, chain'_pair, flip] theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ ∀ (i : ℕ) (h : i < length l - 1), @@ -295,6 +307,10 @@ theorem Chain'.append_overlap {l₁ l₂ l₃ : List α} (h₁ : Chain' R (l₁ h₁.append h₂.right_of_append <| by simpa only [getLast?_append_of_ne_nil _ hn] using (chain'_append.1 h₂).2.2 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → (Chain' R L.join ↔ (∀ l ∈ L, Chain' R l) ∧ L.Chain' (fun l₁ l₂ => ∀ᵉ (x ∈ l₁.getLast?) (y ∈ l₂.head?), R x y)) @@ -304,7 +320,7 @@ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → rw [mem_cons, not_or, ← Ne] at hL rw [join, chain'_append, chain'_join hL.2, forall_mem_cons, chain'_cons] rw [mem_cons, not_or, ← Ne] at hL - simp only [forall_mem_cons, and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] + simp only [forall_mem_cons, _root_.and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] exact Iff.rfl.and (Iff.rfl.and <| Iff.rfl.and and_comm) /-- If `a` and `b` are related by the reflexive transitive closure of `r`, then there is an diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 5f038e2ae1e94..9c8cbaa143209 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -334,7 +334,7 @@ theorem prev_next (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : obtain ⟨⟨n, hn⟩, rfl⟩ := get_of_mem hx simp only [next_get, prev_get, h, Nat.mod_add_mod] cases' l with hd tl - · simp at hx + · simp at hn · have : (n + 1 + length tl) % (length tl + 1) = n := by rw [length_cons] at hn rw [add_assoc, add_comm 1, Nat.add_mod_right, Nat.mod_eq_of_lt hn] @@ -345,7 +345,7 @@ theorem next_prev (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : obtain ⟨⟨n, hn⟩, rfl⟩ := get_of_mem hx simp only [next_get, prev_get, h, Nat.mod_add_mod] cases' l with hd tl - · simp at hx + · simp at hn · have : (n + length tl + 1) % (length tl + 1) = n := by rw [length_cons] at hn rw [add_assoc, Nat.add_mod_right, Nat.mod_eq_of_lt hn] diff --git a/Mathlib/Data/List/Dedup.lean b/Mathlib/Data/List/Dedup.lean index bb2bf2901cd89..55576850f682a 100644 --- a/Mathlib/Data/List/Dedup.lean +++ b/Mathlib/Data/List/Dedup.lean @@ -82,7 +82,7 @@ theorem dedup_eq_cons (l : List α) (a : α) (l' : List α) : l.dedup = a :: l' ↔ a ∈ l ∧ a ∉ l' ∧ l.dedup.tail = l' := by refine ⟨fun h => ?_, fun h => ?_⟩ · refine ⟨mem_dedup.1 (h.symm ▸ mem_cons_self _ _), fun ha => ?_, by rw [h, tail_cons]⟩ - have := count_pos_iff_mem.2 ha + have := count_pos_iff.2 ha have : count a l.dedup ≤ 1 := nodup_iff_count_le_one.1 (nodup_dedup l) a rw [h, count_cons_self] at this omega diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 80750fd2c8b30..f0aad1fcd0f00 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -10,6 +10,7 @@ import Mathlib.Util.CompileInductive import Batteries.Tactic.Lint.Basic import Batteries.Data.List.Lemmas import Batteries.Data.RBMap.Basic +import Batteries.Logic /-! ## Definitions on lists diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 79f5bf6c1cb0a..20808976c60ab 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -84,14 +84,18 @@ 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, and_assoc, + simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, _root_.and_assoc, @and_comm _ (p a), @and_left_comm _ (p a), exists_and_left] - simp only [and_comm, and_assoc, and_left_comm, ← exists_and_right] + simp only [_root_.and_comm, _root_.and_assoc, and_left_comm, ← exists_and_right] @[simp] theorem forall₂_map_left_iff {f : γ → α} : diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index bc144f95b4785..e16ce04be7e33 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -72,16 +72,13 @@ alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq theorem getD_append (l l' : List α) (d : α) (n : ℕ) (h : n < l.length) : (l ++ l').getD n d = l.getD n d := by rw [getD_eq_getElem _ _ (Nat.lt_of_lt_of_le h (length_append _ _ ▸ Nat.le_add_right _ _)), - getElem_append _ h, getD_eq_getElem] + getElem_append_left h, getD_eq_getElem] theorem getD_append_right (l l' : List α) (d : α) (n : ℕ) (h : l.length ≤ n) : (l ++ l').getD n d = l'.getD (n - l.length) d := by cases Nat.lt_or_ge n (l ++ l').length with | inl h' => - rw [getD_eq_getElem (l ++ l') d h', getElem_append_right, getD_eq_getElem] - · rw [length_append] at h' - exact Nat.sub_lt_left_of_lt_add h h' - · exact Nat.not_lt_of_le h + rw [getD_eq_getElem (l ++ l') d h', getElem_append_right h, getD_eq_getElem] | inr h' => rw [getD_eq_default _ _ h', getD_eq_default] rwa [Nat.le_sub_iff_add_le' h, ← length_append] diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index b2351ad435a24..7e53018e89d4b 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -55,9 +55,9 @@ theorem mapIdxGo_append : ∀ (f : ℕ → α → β) (l₁ l₂ : List α) (arr cases l₂ · rfl · rw [List.length_append] at h; contradiction - rw [l₁_nil, l₂_nil]; simp only [mapIdx.go, Array.toList_eq, Array.toArray_data] + rw [l₁_nil, l₂_nil]; simp only [mapIdx.go, List.toArray_toList] · cases' l₁ with head tail <;> simp only [mapIdx.go] - · simp only [nil_append, Array.toList_eq, Array.toArray_data] + · simp only [nil_append, List.toArray_toList] · simp only [List.append_eq] rw [ih] · simp only [cons_append, length_cons, length_append, Nat.succ.injEq] at h @@ -67,7 +67,7 @@ theorem mapIdxGo_length : ∀ (f : ℕ → α → β) (l : List α) (arr : Array length (mapIdx.go f l arr) = length l + arr.size := by intro f l induction' l with head tail ih - · intro; simp only [mapIdx.go, Array.toList_eq, length_nil, Nat.zero_add] + · intro; simp only [mapIdx.go, length_nil, Nat.zero_add] · intro; simp only [mapIdx.go]; rw [ih]; simp only [Array.size_push, length_cons] simp only [Nat.add_succ, Fin.add_zero, Nat.add_comm] @@ -77,7 +77,7 @@ theorem mapIdx_append_one : ∀ (f : ℕ → α → β) (l : List α) (e : α), unfold mapIdx rw [mapIdxGo_append f l [e]] simp only [mapIdx.go, Array.size_toArray, mapIdxGo_length, length_nil, Nat.add_zero, - Array.toList_eq, Array.push_data, Array.data_toArray] + Array.push_toList, Array.toList_toArray] @[local simp] theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α → β), @@ -119,15 +119,16 @@ theorem getElem?_mapIdx_go (f : ℕ → α → β) : ∀ (l : List α) (arr : Ar (mapIdx.go f l arr)[i]? = if h : i < arr.size then some arr[i] else Option.map (f i) l[i - arr.size]? | [], arr, i => by - simp [mapIdx.go, getElem?_eq, Array.getElem_eq_data_getElem] + simp only [mapIdx.go, Array.toListImpl_eq, getElem?_eq, Array.length_toList, + Array.getElem_eq_getElem_toList, length_nil, Nat.not_lt_zero, ↓reduceDIte, Option.map_none'] | a :: l, arr, i => by rw [mapIdx.go, getElem?_mapIdx_go] simp only [Array.size_push] split <;> split · simp only [Option.some.injEq] - rw [Array.getElem_eq_data_getElem] - simp only [Array.push_data] - rw [getElem_append_left, Array.getElem_eq_data_getElem] + rw [Array.getElem_eq_getElem_toList] + simp only [Array.push_toList] + rw [getElem_append_left, Array.getElem_eq_getElem_toList] · have : i = arr.size := by omega simp_all · omega @@ -158,7 +159,7 @@ theorem mapIdx_append (K L : List α) (f : ℕ → α → β) : @[simp] theorem mapIdx_eq_nil {f : ℕ → α → β} {l : List α} : List.mapIdx f l = [] ↔ l = [] := by - rw [List.mapIdx_eq_enum_map, List.map_eq_nil, List.enum_eq_nil] + rw [List.mapIdx_eq_enum_map, List.map_eq_nil_iff, List.enum_eq_nil] theorem get_mapIdx (l : List α) (f : ℕ → α → β) (i : ℕ) (h : i < l.length) (h' : i < (l.mapIdx f).length := h.trans_le (l.length_mapIdx f).ge) : @@ -356,12 +357,12 @@ theorem mapIdxMGo_eq_mapIdxMAuxSpec congr conv => { lhs; intro x; rw [ih _ _ h]; } funext x - simp only [Array.toList_eq, Array.push_data, append_assoc, singleton_append, Array.size_push, + simp only [Array.push_toList, append_assoc, singleton_append, Array.size_push, map_eq_pure_bind] theorem mapIdxM_eq_mmap_enum [LawfulMonad m] {β} (f : ℕ → α → m β) (as : List α) : as.mapIdxM f = List.traverse (uncurry f) (enum as) := by - simp only [mapIdxM, mapIdxMGo_eq_mapIdxMAuxSpec, Array.toList_eq, Array.data_toArray, + simp only [mapIdxM, mapIdxMGo_eq_mapIdxMAuxSpec, Array.toList_toArray, nil_append, mapIdxMAuxSpec, Array.size_toArray, length_nil, id_map', enum] end MapIdxM diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 974beaac0fde7..7e655efa8abee 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -69,8 +69,7 @@ theorem tail_subset (l : List α) : tail l ⊆ l := theorem mem_of_mem_dropLast (h : a ∈ l.dropLast) : a ∈ l := dropLast_subset l h -theorem mem_of_mem_tail (h : a ∈ l.tail) : a ∈ l := - tail_subset l h +attribute [gcongr] Sublist.drop theorem concat_get_prefix {x y : List α} (h : x <+: y) (hl : x.length < y.length) : x ++ [y.get ⟨x.length, hl⟩] <+: y := by @@ -150,7 +149,17 @@ theorem inits_cons (a : α) (l : List α) : inits (a :: l) = [] :: l.inits.map f theorem tails_cons (a : α) (l : List α) : tails (a :: l) = (a :: l) :: l.tails := by simp -@[simp] +#adaptation_note +/-- +This can be removed after nightly-2024-09-07. +-/ +attribute [-simp] map_tail + +#adaptation_note +/-- +`nolint simpNF` should be removed after nightly-2024-09-07. +-/ +@[simp, nolint simpNF] theorem inits_append : ∀ s t : List α, inits (s ++ t) = s.inits ++ t.inits.tail.map fun l => s ++ l | [], [] => by simp | [], a :: t => by simp diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index 99646ea33a4c5..b6b147595d8d8 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -83,13 +83,17 @@ theorem insertNth_comm (a b : α) : 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 | n + 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), - ← or_assoc, @or_comm (a = a'), or_assoc, mem_cons] + ← _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 diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index c1419e19f2901..870f434107df2 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -135,7 +135,7 @@ theorem filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : simp only [(lt_of_lt_of_le (mem.1 hk).2 hml), decide_True] theorem filter_lt_of_le_bot {n m l : ℕ} (hln : l ≤ n) : ((Ico n m).filter fun x => x < l) = [] := - filter_eq_nil.2 fun k hk => by + filter_eq_nil_iff.2 fun k hk => by simp only [decide_eq_true_eq, not_lt] apply le_trans hln exact (mem.1 hk).1 @@ -161,7 +161,7 @@ theorem filter_le_of_le_bot {n m l : ℕ} (hln : l ≤ n) : exact le_trans hln (mem.1 hk).1 theorem filter_le_of_top_le {n m l : ℕ} (hml : m ≤ l) : ((Ico n m).filter fun x => l ≤ x) = [] := - filter_eq_nil.2 fun k hk => by + filter_eq_nil_iff.2 fun k hk => by rw [decide_eq_true_eq] exact not_le_of_gt (lt_of_lt_of_le (mem.1 hk).2 hml) diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 7c3ce41155f9f..75dc0c1e112bf 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -203,7 +203,7 @@ theorem count_bagInter {a : α} : by_cases ba : b = a · simp only [beq_iff_eq] rw [if_pos ba, Nat.sub_add_cancel] - rwa [succ_le_iff, count_pos_iff_mem, ← ba] + rwa [succ_le_iff, count_pos_iff, ← ba] · simp only [beq_iff_eq] rw [if_neg ba, Nat.sub_zero, Nat.add_zero, Nat.add_zero] · rw [cons_bagInter_of_neg _ hb, count_bagInter] diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index ea2d1faa23b35..b055f1b187222 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -30,14 +30,6 @@ theorem tail_reverse_eq_reverse_dropLast (l : List α) : · rw [getElem?_eq_none, getElem?_eq_none] all_goals (simp; omega) -theorem getLast_tail (l : List α) (hl : l.tail ≠ []) : - l.tail.getLast hl = l.getLast (by intro h; rw [h] at hl; simp at hl) := by - simp only [← drop_one, ne_eq, drop_eq_nil_iff_le, - not_le, getLast_eq_getElem, length_drop] at hl |- - rw [← getElem_drop'] - · simp [show 1 + (l.length - 1 - 1) = l.length - 1 by omega] - omega - @[deprecated (since := "2024-08-19")] alias nthLe_tail := getElem_tail theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index c20935506c371..52a8f59fb7167 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -433,25 +433,31 @@ theorem minimum_of_length_pos_le_getElem {i : ℕ} (w : i < l.length) (h := (Nat l.minimum_of_length_pos h ≤ l[i] := getElem_le_maximum_of_length_pos (α := αᵒᵈ) w -lemma getD_maximum?_eq_unbot'_maximum (l : List α) (d : α) : - l.maximum?.getD d = l.maximum.unbot' d := by +lemma getD_max?_eq_unbot'_maximum (l : List α) (d : α) : + l.max?.getD d = l.maximum.unbot' d := by cases hy : l.maximum with | bot => simp [List.maximum_eq_bot.mp hy] | coe y => rw [List.maximum_eq_coe_iff] at hy simp only [WithBot.unbot'_coe] - cases hz : l.maximum? with - | none => simp [List.maximum?_eq_none_iff.mp hz] at hy + cases hz : l.max? with + | none => simp [List.max?_eq_none_iff.mp hz] at hy | some z => have : Antisymm (α := α) (· ≤ ·) := ⟨_root_.le_antisymm⟩ - rw [List.maximum?_eq_some_iff] at hz + rw [List.max?_eq_some_iff] at hz · rw [Option.getD_some] exact _root_.le_antisymm (hy.right _ hz.left) (hz.right _ hy.left) all_goals simp [le_total] -lemma getD_minimum?_eq_untop'_minimum (l : List α) (d : α) : - l.minimum?.getD d = l.minimum.untop' d := - getD_maximum?_eq_unbot'_maximum (α := αᵒᵈ) _ _ +@[deprecated (since := "2024-09-29")] +alias getD_maximum?_eq_unbot'_maximum := getD_max?_eq_unbot'_maximum + +lemma getD_min?_eq_untop'_minimum (l : List α) (d : α) : + l.min?.getD d = l.minimum.untop' d := + getD_max?_eq_unbot'_maximum (α := αᵒᵈ) _ _ + +@[deprecated (since := "2024-09-29")] +alias getD_minimum?_eq_untop'_minimum := getD_min?_eq_untop'_minimum end LinearOrder diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index c4ede7b589049..9d86abe857e93 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -16,7 +16,7 @@ predicate. universe u v -open Nat Function +open Function variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a b : α} @@ -102,13 +102,17 @@ theorem nodup_iff_get?_ne_get? {l : List α} : l.Nodup ↔ ∀ i j : ℕ, i < j → j < l.length → l.get? i ≠ l.get? j := by simp [nodup_iff_getElem?_ne_getElem?] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem Nodup.ne_singleton_iff {l : List α} (h : Nodup l) (x : α) : l ≠ [x] ↔ l = [] ∨ ∃ y ∈ l, y ≠ x := by induction' l with hd tl hl · simp · specialize hl h.of_cons by_cases hx : tl = [x] - · simpa [hx, and_comm, and_or_left] using h + · simpa [hx, _root_.and_comm, and_or_left] using h · rw [← Ne, hl] at hx rcases hx with (rfl | ⟨y, hy, hx⟩) · simp @@ -140,14 +144,14 @@ theorem nodup_iff_count_le_one [DecidableEq α] {l : List α} : Nodup l ↔ ∀ theorem nodup_iff_count_eq_one [DecidableEq α] : Nodup l ↔ ∀ a ∈ l, count a l = 1 := nodup_iff_count_le_one.trans <| forall_congr' fun _ => - ⟨fun H h => H.antisymm (count_pos_iff_mem.mpr h), + ⟨fun H h => H.antisymm (count_pos_iff.mpr h), fun H => if h : _ then (H h).le else (count_eq_zero.mpr h).trans_le (Nat.zero_le 1)⟩ @[simp] theorem count_eq_one_of_mem [DecidableEq α] {a : α} {l : List α} (d : Nodup l) (h : a ∈ l) : count a l = 1 := - _root_.le_antisymm (nodup_iff_count_le_one.1 d a) (Nat.succ_le_of_lt (count_pos_iff_mem.2 h)) + _root_.le_antisymm (nodup_iff_count_le_one.1 d a) (Nat.succ_le_of_lt (count_pos_iff.2 h)) theorem count_eq_of_nodup [DecidableEq α] {a : α} {l : List α} (d : Nodup l) : count a l = if a ∈ l then 1 else 0 := by @@ -174,9 +178,13 @@ theorem Nodup.append (d₁ : Nodup l₁) (d₂ : Nodup l₂) (dj : Disjoint l₁ theorem nodup_append_comm {l₁ l₂ : List α} : Nodup (l₁ ++ l₂) ↔ Nodup (l₂ ++ l₁) := by simp only [nodup_append, and_left_comm, disjoint_comm] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem nodup_middle {a : α} {l₁ l₂ : List α} : Nodup (l₁ ++ a :: l₂) ↔ Nodup (a :: (l₁ ++ l₂)) := by - simp only [nodup_append, not_or, and_left_comm, and_assoc, nodup_cons, mem_append, + simp only [nodup_append, not_or, and_left_comm, _root_.and_assoc, nodup_cons, mem_append, disjoint_cons_right] theorem Nodup.of_map (f : α → β) {l : List α} : Nodup (map f l) → Nodup l := @@ -244,8 +252,8 @@ theorem Nodup.erase_getElem [DecidableEq α] {l : List α} (hl : l.Nodup) · simp [IH hl.2] · rw [beq_iff_eq] simp only [getElem_cons_succ] - simp only [length_cons, succ_eq_add_one, Nat.add_lt_add_iff_right] at h - exact mt (· ▸ l.getElem_mem i h) hl.1 + simp only [length_cons, Nat.succ_eq_add_one, Nat.add_lt_add_iff_right] at h + exact mt (· ▸ getElem_mem h) hl.1 theorem Nodup.erase_get [DecidableEq α] {l : List α} (hl : l.Nodup) (i : Fin l.length) : l.erase (l.get i) = l.eraseIdx ↑i := by @@ -259,11 +267,15 @@ theorem nodup_join {L : List (List α)} : Nodup (join L) ↔ (∀ l ∈ L, Nodup l) ∧ Pairwise Disjoint L := by simp only [Nodup, pairwise_join, disjoint_left.symm, forall_mem_ne] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem nodup_bind {l₁ : List α} {f : α → List β} : Nodup (l₁.bind f) ↔ (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (fun a b : α => Disjoint (f a) (f b)) l₁ := by - simp only [List.bind, nodup_join, pairwise_map, and_comm, and_left_comm, mem_map, exists_imp, - and_imp] + simp only [List.bind, nodup_join, pairwise_map, _root_.and_comm, and_left_comm, mem_map, + exists_imp, and_imp] rw [show (∀ (l : List β) (x : α), f x = l → x ∈ l₁ → Nodup l) ↔ ∀ x : α, x ∈ l₁ → Nodup (f x) from forall_swap.trans <| forall_congr' fun _ => forall_eq'] @@ -304,13 +316,12 @@ theorem Nodup.union [DecidableEq α] (l₁ : List α) (h : Nodup l₂) : (l₁ theorem Nodup.inter [DecidableEq α] (l₂ : List α) : Nodup l₁ → Nodup (l₁ ∩ l₂) := Nodup.filter _ -theorem Nodup.diff_eq_filter [DecidableEq α] : +theorem Nodup.diff_eq_filter [BEq α] [LawfulBEq α] : ∀ {l₁ l₂ : List α} (_ : l₁.Nodup), l₁.diff l₂ = l₁.filter (· ∉ l₂) | l₁, [], _ => by simp | l₁, a :: l₂, hl₁ => by rw [diff_cons, (hl₁.erase _).diff_eq_filter, hl₁.erase_eq_filter, filter_filter] - simp only [decide_not, Bool.not_eq_true', decide_eq_false_iff_not, bne_iff_ne, ne_eq, and_comm, - Bool.decide_and, mem_cons, not_or] + simp only [decide_not, bne, Bool.and_comm, mem_cons, not_or, decide_mem_cons, Bool.not_or] theorem Nodup.mem_diff_iff [DecidableEq α] (hl₁ : l₁.Nodup) : a ∈ l₁.diff l₂ ↔ a ∈ l₁ ∧ a ∉ l₂ := by rw [hl₁.diff_eq_filter, mem_filter, decide_eq_true_iff] diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 5266ae7b60164..0dcf11280caae 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -126,7 +126,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) exact ix.succ_pos rw [← List.take_append_drop (f 0 + 1) l', ← List.singleton_append] apply List.Sublist.append _ (IH _ this) - rw [List.singleton_sublist, ← h, l'.getElem_take _ (Nat.lt_succ_self _)] + rw [List.singleton_sublist, ← h, l'.getElem_take' _ (Nat.lt_succ_self _)] apply List.get_mem /-- A `l : List α` is `Sublist l l'` for `l' : List α` iff diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index fdff37f5163e5..f0e74dcc0d2b3 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -51,25 +51,6 @@ theorem Perm.subset_congr_left {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l theorem Perm.subset_congr_right {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l₃ ⊆ l₁ ↔ l₃ ⊆ l₂ := ⟨fun h' => h'.trans h.subset, fun h' => h'.trans h.symm.subset⟩ -/-- Variant of `Perm.foldr_eq` with explicit commutativity argument. -/ -theorem Perm.foldr_eq' {f : α → β → β} {l₁ l₂ : List α} (p : l₁ ~ l₂) - (comm : ∀ x ∈ l₁, ∀ y ∈ l₁, ∀ z, f y (f x z) = f x (f y z)) - (init : β) : foldr f init l₁ = foldr f init l₂ := by - induction p using recOnSwap' generalizing init with - | nil => simp - | cons x _p IH => - simp only [foldr] - congr 1 - apply IH; intros; apply comm <;> exact .tail _ ‹_› - | swap' x y _p IH => - simp only [foldr] - rw [comm x (.tail _ <| .head _) y (.head _)] - congr 2 - apply IH; intros; apply comm <;> exact .tail _ (.tail _ ‹_›) - | trans p₁ _p₂ IH₁ IH₂ => - refine (IH₁ comm init).trans (IH₂ ?_ _) - intros; apply comm <;> apply p₁.symm.subset <;> assumption - section Rel open Relator @@ -265,13 +246,17 @@ theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t l₁.bagInter t₁ ~ l₂.bagInter t₂ := ht.bagInter_left l₂ ▸ hl.bagInter_right _ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a - · simp [subset_def, or_comm] + · simp [subset_def, _root_.or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index dfa4674956933..37fb13a041412 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -166,6 +166,10 @@ theorem foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : · rfl · simp_rw [foldr_cons, ih, bind_cons, append_assoc, permutationsAux2_append] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α)} {l' : List α} : l' ∈ foldr (fun y r => (permutationsAux2 t ts r y id).2) r L ↔ l' ∈ r ∨ ∃ l₁ l₂, l₁ ++ l₂ ∈ L ∧ l₂ ≠ [] ∧ l' = l₁ ++ t :: l₂ ++ ts := by @@ -176,7 +180,7 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) ⟨fun ⟨_, aL, l₁, l₂, l0, e, h⟩ => ⟨l₁, l₂, l0, e ▸ aL, h⟩, fun ⟨l₁, l₂, l0, aL, h⟩ => ⟨_, aL, l₁, l₂, l0, rfl, h⟩⟩ rw [foldr_permutationsAux2] - simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_bind, + simp only [mem_permutationsAux2', ← this, _root_.or_comm, and_left_comm, mem_append, mem_bind, append_assoc, cons_append, exists_prop] theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 5cbd6337a5630..923a8ba69b0ed 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -104,11 +104,15 @@ theorem rotate_cons_succ (l : List α) (a : α) (n : ℕ) : (a :: l : List α).rotate (n + 1) = (l ++ [a]).rotate n := by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_rotate : ∀ {l : List α} {a : α} {n : ℕ}, a ∈ l.rotate n ↔ a ∈ l | [], _, n => by simp | a :: l, _, 0 => by simp - | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, or_comm] + | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, _root_.or_comm] @[simp] theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length := by @@ -116,7 +120,7 @@ theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length : @[simp] theorem rotate_replicate (a : α) (n : ℕ) (k : ℕ) : (replicate n a).rotate k = replicate n a := - eq_replicate.2 ⟨by rw [length_rotate, length_replicate], fun b hb => + eq_replicate_iff.2 ⟨by rw [length_rotate, length_replicate], fun b hb => eq_of_mem_replicate <| mem_rotate.1 hb⟩ theorem rotate_eq_drop_append_take {l : List α} {n : ℕ} : @@ -473,7 +477,7 @@ theorem IsRotated.dropLast_tail {α} | [] => by simp | [_] => by simp | a :: b :: L => by - simp at hL' |- + simp only [head_cons, ne_eq, reduceCtorEq, not_false_eq_true, getLast_cons] at hL' simp [hL', IsRotated.cons_getLast_dropLast] /-- List of all cyclic permutations of `l`. diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 9742e4370e93c..162ca0b8a4138 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -123,8 +123,8 @@ theorem eq_of_perm_of_sorted [IsAntisymm α r] {l₁ l₂ : List α} (hp : l₁ congr have : ∀ x ∈ u₂, x = a := fun x m => antisymm ((pairwise_append.1 hs₂).2.2 _ m a (mem_cons_self _ _)) (h₁ _ (by simp [m])) - rw [(@eq_replicate _ a (length u₂ + 1) (a :: u₂)).2, - (@eq_replicate _ a (length u₂ + 1) (u₂ ++ [a])).2] <;> + rw [(@eq_replicate_iff _ a (length u₂ + 1) (a :: u₂)).2, + (@eq_replicate_iff _ a (length u₂ + 1) (u₂ ++ [a])).2] <;> constructor <;> simp [iff_true_intro this, or_comm] @@ -149,7 +149,7 @@ theorem Sorted.rel_of_mem_take_of_mem_drop {l : List α} (h : List.Sorted r l) { (hx : x ∈ List.take k l) (hy : y ∈ List.drop k l) : r x y := by obtain ⟨iy, hiy, rfl⟩ := getElem_of_mem hy obtain ⟨ix, hix, rfl⟩ := getElem_of_mem hx - rw [getElem_take', getElem_drop] + rw [getElem_take, getElem_drop] rw [length_take] at hix exact h.rel_get_of_lt (Nat.lt_add_right _ (Nat.lt_min.mp hix).left) @@ -534,12 +534,12 @@ def mergeSort' : List α → List α let ls := (split (a :: b :: l)) have := length_split_fst_le l have := length_split_snd_le l - exact merge (r · ·) (mergeSort' ls.1) (mergeSort' ls.2) + 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 (r · ·) (mergeSort' r l₁) (mergeSort' r l₂) := by + mergeSort' r (a :: b :: l) = merge (mergeSort' r l₁) (mergeSort' r l₂) (r · ·) := by simp only [mergeSort', h] section Correctness @@ -568,13 +568,13 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge (r · ·) l l') +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 (r · ·) l (b :: l'), r a b' by + · 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 @@ -584,7 +584,7 @@ theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sort assumption · exact rel_of_sorted_cons h₁ _ bl · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge (r · ·) (a :: l) l', r b b' by + · 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 @@ -625,24 +625,6 @@ theorem mergeSort'_nil : [].mergeSort' r = [] := by rw [List.mergeSort'] @[simp] theorem mergeSort'_singleton (a : α) : [a].mergeSort' r = [a] := by rw [List.mergeSort'] -theorem map_merge (f : α → β) (r : α → α → Bool) (s : β → β → Bool) (l l' : List α) - (hl : ∀ a ∈ l, ∀ b ∈ l', r a b = s (f a) (f b)) : - (l.merge r l').map f = (l.map f).merge s (l'.map f) := by - match l, l' with - | [], x' => simp - | x, [] => simp - | x :: xs, x' :: xs' => - simp_rw [List.forall_mem_cons, forall_and] at hl - simp_rw [List.map, List.cons_merge_cons] - rw [← hl.1.1] - split - · rw [List.map, map_merge _ r s, List.map] - simp_rw [List.forall_mem_cons, forall_and] - exact ⟨hl.2.1, hl.2.2⟩ - · rw [List.map, map_merge _ r s, List.map] - simp_rw [List.forall_mem_cons] - exact ⟨hl.1.2, hl.2.2⟩ - 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 @@ -659,7 +641,7 @@ theorem map_mergeSort' (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b have := length_split_snd_le l simp_rw [List.map] rw [List.mergeSort'_cons_cons _ e, List.mergeSort'_cons_cons _ fe, - map_merge _ (r · ·) (s · ·), map_mergeSort' _ l₁ hl.1.1, map_mergeSort' _ l₂ hl.2.2] + 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 diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 8eac342a35b41..3d8300bd2d98a 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -44,7 +44,7 @@ theorem sublists'Aux_eq_array_foldl (a : α) : ∀ (r₁ r₂ : List (List α)), sublists'Aux a r₁ r₂ = ((r₁.toArray).foldl (init := r₂.toArray) (fun r l => r.push (a :: l))).toList := by intro r₁ r₂ - rw [sublists'Aux, Array.foldl_eq_foldl_data] + rw [sublists'Aux, Array.foldl_eq_foldl_toList] have := List.foldl_hom Array.toList (fun r l => r.push (a :: l)) (fun r l => r ++ [a :: l]) r₁ r₂.toArray (by simp) simpa using this @@ -53,8 +53,7 @@ theorem sublists'_eq_sublists'Aux (l : List α) : sublists' l = l.foldr (fun a r => sublists'Aux a r r) [[]] := by simp only [sublists', sublists'Aux_eq_array_foldl] rw [← List.foldr_hom Array.toList] - · rfl - · intros _ _; congr <;> simp + · intros _ _; congr theorem sublists'Aux_eq_map (a : α) (r₁ : List (List α)) : ∀ (r₂ : List (List α)), sublists'Aux a r₁ r₂ = r₂ ++ map (cons a) r₁ := @@ -107,7 +106,7 @@ theorem sublistsAux_eq_array_foldl : (r.toArray.foldl (init := #[]) fun r l => (r.push l).push (a :: l)).toList := by funext a r - simp only [sublistsAux, Array.foldl_eq_foldl_data, Array.mkEmpty] + simp only [sublistsAux, Array.foldl_eq_foldl_toList, Array.mkEmpty] have := foldl_hom Array.toList (fun r l => (r.push l).push (a :: l)) (fun (r : List (List α)) l => r ++ [l, a :: l]) r #[] (by simp) @@ -126,10 +125,9 @@ theorem sublistsAux_eq_bind : ext α l : 2 trans l.foldr sublistsAux [[]] · rw [sublistsAux_eq_bind, sublists] - · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_data] + · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_toList] rw [← foldr_hom Array.toList] - · rfl - · intros _ _; congr <;> simp + · intros _ _; congr theorem sublists_append (l₁ l₂ : List α) : sublists (l₁ ++ l₂) = (sublists l₂) >>= (fun x => (sublists l₁).map (· ++ x)) := by diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index 80954c1ad9dc7..1d6228b9258e3 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -237,7 +237,7 @@ theorem sym_one_eq : xs.sym 1 = xs.map (· ::ₛ .nil) := by theorem sym2_eq_sym_two : xs.sym2.map (Sym2.equivSym α) = xs.sym 2 := by induction xs with - | nil => simp only [List.sym, map_eq_nil, sym2_eq_nil_iff] + | nil => simp only [List.sym, map_eq_nil_iff, sym2_eq_nil_iff] | cons x xs ih => rw [List.sym, ← ih, sym_one_eq, map_map, List.sym2, map_append, map_map] rfl diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index c124885555f3b..6991b7dc20d72 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1345,8 +1345,7 @@ theorem pmap_eq_map (p : α → Prop) (f : α → β) (s : Multiset α) : theorem pmap_congr {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (s : Multiset α) : ∀ {H₁ H₂}, (∀ a ∈ s, ∀ (h₁ h₂), f a h₁ = g a h₂) → pmap f s H₁ = pmap g s H₂ := - @(Quot.inductionOn s (fun l _H₁ _H₂ h => congr_arg _ <| List.pmap_congr l h)) - + @(Quot.inductionOn s (fun l _H₁ _H₂ h => congr_arg _ <| List.pmap_congr_left l h)) theorem map_pmap {p : α → Prop} (g : β → γ) (f : ∀ a, p a → β) (s) : ∀ H, map g (pmap f s H) = pmap (fun a h => g (f a h)) s H := @@ -1391,7 +1390,7 @@ theorem attach_cons (a : α) (m : Multiset α) : Quotient.inductionOn m fun l => congr_arg _ <| congr_arg (List.cons _) <| by - rw [List.map_pmap]; exact List.pmap_congr _ fun _ _ _ _ => Subtype.eq rfl + rw [List.map_pmap]; exact List.pmap_congr_left _ fun _ _ _ _ => Subtype.eq rfl section DecidablePiExists @@ -1828,17 +1827,25 @@ theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := Quot.inductionOn s fun l => by simp +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma filter_comm (q) [DecidablePred q] (s : Multiset α) : - filter p (filter q s) = filter q (filter p s) := by simp [and_comm] + filter p (filter q s) = filter q (filter p s) := by simp [_root_.and_comm] theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) : filter p s + filter q s = filter (fun a => p a ∨ q a) s + filter (fun a => p a ∧ q a) s := Multiset.induction_on s rfl fun a s IH => by by_cases p a <;> by_cases q a <;> simp [*] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) s = s := by rw [filter_add_filter, filter_eq_self.2, filter_eq_nil.2] · simp only [add_zero] - · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, or_comm] + · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, _root_.or_comm] · simp only [Bool.not_eq_true, decide_eq_true_eq, Bool.eq_false_or_eq_true, decide_True, implies_true, Decidable.em] @@ -2017,11 +2024,11 @@ theorem countP_eq_countP_filter_add (s) (p q : α → Prop) [DecidablePred p] [D @[simp] theorem countP_True {s : Multiset α} : countP (fun _ => True) s = card s := - Quot.inductionOn s fun _l => List.countP_true + Quot.inductionOn s fun _l => congrFun List.countP_true _ @[simp] theorem countP_False {s : Multiset α} : countP (fun _ => False) s = 0 := - Quot.inductionOn s fun _l => List.countP_false + Quot.inductionOn s fun _l => congrFun List.countP_false _ theorem countP_map (f : α → β) (s : Multiset α) (p : β → Prop) [DecidablePred p] : countP p (map f s) = card (s.filter fun a => p (f a)) := by @@ -2050,7 +2057,7 @@ lemma filter_attach (s : Multiset α) (p : α → Prop) [DecidablePred p] : variable {p} theorem countP_pos {s} : 0 < countP p s ↔ ∃ a ∈ s, p a := - Quot.inductionOn s fun _l => by simpa using List.countP_pos (p ·) + Quot.inductionOn s fun _l => by simp theorem countP_eq_zero {s} : countP p s = 0 ↔ ∀ a ∈ s, ¬p a := Quot.inductionOn s fun _l => by simp [List.countP_eq_zero] @@ -2678,9 +2685,13 @@ lemma add_eq_union_left_of_le [DecidableEq α] {s t u : Multiset α} (h : t ≤ · rintro ⟨h0, rfl⟩ exact h0 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma add_eq_union_right_of_le [DecidableEq α] {x y z : Multiset α} (h : z ≤ y) : x + y = x ∪ z ↔ y = z ∧ x.Disjoint y := by - simpa only [and_comm] using add_eq_union_left_of_le h + simpa only [_root_.and_comm] using add_eq_union_left_of_le h theorem disjoint_map_map {f : α → γ} {g : β → γ} {s : Multiset α} {t : Multiset β} : Disjoint (s.map f) (t.map g) ↔ ∀ a ∈ s, ∀ b ∈ t, f a ≠ g b := by diff --git a/Mathlib/Data/Multiset/Functor.lean b/Mathlib/Data/Multiset/Functor.lean index 9c8f9ecbfb459..81575f5f445c4 100644 --- a/Mathlib/Data/Multiset/Functor.lean +++ b/Mathlib/Data/Multiset/Functor.lean @@ -99,7 +99,6 @@ theorem comp_traverse {G H : Type _ → Type _} [Applicative G] [Applicative H] intro simp only [traverse, quot_mk_to_coe, lift_coe, Coe.coe, Function.comp_apply, Functor.map_map, functor_norm] - simp only [Function.comp_def, lift_coe] theorem map_traverse {G : Type* → Type _} [Applicative G] [CommApplicative G] {α β γ : Type _} (g : α → G β) (h : β → γ) (x : Multiset α) : @@ -107,7 +106,8 @@ theorem map_traverse {G : Type* → Type _} [Applicative G] [CommApplicative G] refine Quotient.inductionOn x ?_ intro simp only [traverse, quot_mk_to_coe, lift_coe, Function.comp_apply, Functor.map_map, map_comp_coe] - rw [LawfulFunctor.comp_map, Traversable.map_traverse'] + rw [Traversable.map_traverse'] + simp only [fmap_def, Function.comp_apply, Functor.map_map, List.map_eq_map] rfl theorem traverse_map {G : Type* → Type _} [Applicative G] [CommApplicative G] {α β γ : Type _} diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 8e257d61ead92..79bd2eb583ef4 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -35,9 +35,9 @@ elements of `s` in increasing order. -/ def bitIndices (n : ℕ) : List ℕ := @binaryRec (fun _ ↦ List ℕ) [] (fun b _ s ↦ b.casesOn (s.map (· + 1)) (0 :: s.map (· + 1))) n -@[simp] theorem bitIndices_zero : bitIndices 0 = [] := by rfl +@[simp] theorem bitIndices_zero : bitIndices 0 = [] := by simp [bitIndices] -@[simp] theorem bitIndices_one : bitIndices 1 = [0] := by rfl +@[simp] theorem bitIndices_one : bitIndices 1 = [0] := by simp [bitIndices] theorem bitIndices_bit_true (n : ℕ) : bitIndices (bit true n) = 0 :: ((bitIndices n).map (· + 1)) := diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index feb4a5b86056f..753381e7318d1 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -25,7 +25,7 @@ and `Nat.digits`. -- Once we're in the `Nat` namespace, `xor` will inconveniently resolve to `Nat.xor`. /-- `bxor` denotes the `xor` function i.e. the exclusive-or function on type `Bool`. -/ -local notation "bxor" => _root_.xor +local notation "bxor" => xor namespace Nat universe u @@ -48,7 +48,7 @@ def bodd (n : ℕ) : Bool := (boddDiv2 n).1 @[simp] lemma bodd_zero : bodd 0 = false := rfl -lemma bodd_one : bodd 1 = true := rfl +@[simp] lemma bodd_one : bodd 1 = true := rfl lemma bodd_two : bodd 2 = false := rfl @@ -88,12 +88,12 @@ lemma mod_two_of_bodd (n : ℕ) : n % 2 = cond (bodd n) 1 0 := by @[simp] lemma div2_zero : div2 0 = 0 := rfl -lemma div2_one : div2 1 = 0 := rfl +@[simp] lemma div2_one : div2 1 = 0 := rfl lemma div2_two : div2 2 = 1 := rfl @[simp] -lemma div2_succ (n : ℕ) : div2 (succ n) = cond (bodd n) (succ (div2 n)) (div2 n) := by +lemma div2_succ (n : ℕ) : div2 (n + 1) = cond (bodd n) (succ (div2 n)) (div2 n) := by simp only [bodd, boddDiv2, div2] rcases boddDiv2 n with ⟨_|_, _⟩ <;> simp @@ -195,6 +195,12 @@ lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit rw [binaryRec] rfl +@[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 + rw [binaryRec] + simp + /-! bitwise ops -/ lemma bodd_bit (b n) : bodd (bit b n) = b := by @@ -391,6 +397,7 @@ theorem bit1_bits (n : ℕ) : (2 * n + 1).bits = true :: n.bits := @[simp] theorem one_bits : Nat.bits 1 = [true] := by convert bit1_bits 0 + simp -- TODO Find somewhere this can live. -- example : bits 3423 = [true, true, true, true, true, false, true, false, true, false, true, true] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index ea2eb8e119049..98bc466c54c8d 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -9,6 +9,7 @@ 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 @@ -272,9 +273,6 @@ theorem lor_comm (n m : ℕ) : n ||| m = m ||| n := theorem land_comm (n m : ℕ) : n &&& m = m &&& n := bitwise_comm Bool.and_comm n m -protected lemma xor_comm (n m : ℕ) : n ^^^ m = m ^^^ n := - bitwise_comm (Bool.bne_eq_xor ▸ Bool.xor_comm) n m - lemma and_two_pow (n i : ℕ) : n &&& 2 ^ i = (n.testBit i).toNat * 2 ^ i := by refine eq_of_testBit_eq fun j => ?_ obtain rfl | hij := Decidable.eq_or_ne i j <;> cases' h : n.testBit i @@ -286,13 +284,6 @@ lemma and_two_pow (n i : ℕ) : n &&& 2 ^ i = (n.testBit i).toNat * 2 ^ i := by lemma two_pow_and (n i : ℕ) : 2 ^ i &&& n = 2 ^ i * (n.testBit i).toNat := by rw [mul_comm, land_comm, and_two_pow] -@[simp] -theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by simp [HXor.hXor, Xor.xor, xor] - -@[simp] -theorem xor_zero (n : ℕ) : n ^^^ 0 = n := by simp [HXor.hXor, Xor.xor, xor] - - /-- Proving associativity of bitwise operations in general essentially boils down to a huge case distinction, so it is shorter to use this tactic instead of proving it in the general case. -/ macro "bitwise_assoc_tac" : tactic => set_option hygiene false in `(tactic| ( @@ -305,22 +296,16 @@ macro "bitwise_assoc_tac" : tactic => set_option hygiene false in `(tactic| ( -- This is necessary because these are simp lemmas in mathlib <;> simp [hn, Bool.or_assoc, Bool.and_assoc, Bool.bne_eq_xor])) -protected lemma xor_assoc (n m k : ℕ) : (n ^^^ m) ^^^ k = n ^^^ (m ^^^ k) := by bitwise_assoc_tac - theorem land_assoc (n m k : ℕ) : (n &&& m) &&& k = n &&& (m &&& k) := by bitwise_assoc_tac theorem lor_assoc (n m k : ℕ) : (n ||| m) ||| k = n ||| (m ||| k) := by bitwise_assoc_tac -@[simp] -theorem xor_self (n : ℕ) : n ^^^ n = 0 := - zero_of_testBit_eq_false fun i => by simp - -- These lemmas match `mul_inv_cancel_right` and `mul_inv_cancel_left`. theorem xor_cancel_right (n m : ℕ) : (m ^^^ n) ^^^ n = m := by - rw [Nat.xor_assoc, xor_self, xor_zero] + rw [Nat.xor_assoc, Nat.xor_self, xor_zero] theorem xor_cancel_left (n m : ℕ) : n ^^^ (n ^^^ m) = m := by - rw [← Nat.xor_assoc, xor_self, zero_xor] + rw [← Nat.xor_assoc, Nat.xor_self, zero_xor] theorem xor_right_injective {n : ℕ} : Function.Injective (HXor.hXor n : ℕ → ℕ) := fun m m' h => by rw [← xor_cancel_left n m, ← xor_cancel_left n m', h] @@ -339,7 +324,7 @@ theorem xor_left_inj {n m m' : ℕ} : m ^^^ n = m' ^^^ n ↔ m = m' := @[simp] theorem xor_eq_zero {n m : ℕ} : n ^^^ m = 0 ↔ n = m := by - rw [← xor_self n, xor_right_inj, eq_comm] + rw [← Nat.xor_self n, xor_right_inj, eq_comm] theorem xor_ne_zero {n m : ℕ} : n ^^^ m ≠ 0 ↔ n ≠ m := xor_eq_zero.not diff --git a/Mathlib/Data/Nat/Cast/NeZero.lean b/Mathlib/Data/Nat/Cast/NeZero.lean index d49f78b3b3095..c3f00eb8b8820 100644 --- a/Mathlib/Data/Nat/Cast/NeZero.lean +++ b/Mathlib/Data/Nat/Cast/NeZero.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Gabriel Ebner -/ import Mathlib.Data.Nat.Cast.Defs -import Mathlib.Algebra.NeZero /-! # Lemmas about nonzero elements of an `AddMonoidWithOne` diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 81287ae2f16f5..4ee7eb93975ba 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -137,8 +137,6 @@ lemma one_lt_iff_ne_zero_and_ne_one : ∀ {n : ℕ}, 1 < n ↔ n ≠ 0 ∧ n ≠ lemma le_one_iff_eq_zero_or_eq_one : ∀ {n : ℕ}, n ≤ 1 ↔ n = 0 ∨ n = 1 := by simp [le_succ_iff] -@[simp] lemma lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans <| by rw [le_zero_eq] - lemma one_le_of_lt (h : a < b) : 1 ≤ b := Nat.lt_of_le_of_lt (Nat.zero_le _) h protected lemma min_left_comm (a b c : ℕ) : min a (min b c) = min b (min a c) := by @@ -167,8 +165,13 @@ 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, or_comm, forall_eq_or_imp, and_comm] + simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, _root_.or_comm, forall_eq_or_imp, + _root_.and_comm] lemma exists_lt_succ : (∃ m < n + 1, p m) ↔ (∃ m < n, p m) ∨ p n := by rw [← not_iff_not] @@ -310,11 +313,11 @@ lemma two_mul_ne_two_mul_add_one : 2 * n ≠ 2 * m + 1 := -- TODO: Replace `Nat.mul_right_cancel_iff` with `Nat.mul_left_inj` protected lemma mul_left_inj (ha : a ≠ 0) : b * a = c * a ↔ b = c := - Nat.mul_right_cancel_iff (Nat.pos_iff_ne_zero.2 ha) _ _ + Nat.mul_right_cancel_iff (Nat.pos_iff_ne_zero.2 ha) -- TODO: Replace `Nat.mul_left_cancel_iff` with `Nat.mul_right_inj` protected lemma mul_right_inj (ha : a ≠ 0) : a * b = a * c ↔ b = c := - Nat.mul_left_cancel_iff (Nat.pos_iff_ne_zero.2 ha) _ _ + Nat.mul_left_cancel_iff (Nat.pos_iff_ne_zero.2 ha) protected lemma mul_ne_mul_left (ha : a ≠ 0) : b * a ≠ c * a ↔ b ≠ c := not_congr (Nat.mul_left_inj ha) @@ -834,7 +837,7 @@ This is an alias of `Nat.leRec`, specialized to `Prop`. -/ @[elab_as_elim] lemma le_induction {m : ℕ} {P : ∀ n, m ≤ n → Prop} (base : P m m.le_refl) (succ : ∀ n hmn, P n hmn → P (n + 1) (le_succ_of_le hmn)) : ∀ n hmn, P n hmn := - @Nat.leRec (motive := P) base succ + @Nat.leRec (motive := P) _ base succ /-- Induction principle deriving the next case from the two previous ones. -/ def twoStepInduction {P : ℕ → Sort*} (zero : P 0) (one : P 1) @@ -1018,9 +1021,6 @@ lemma div_ne_zero_iff_of_dvd (hba : b ∣ a) : a / b ≠ 0 ↔ a ≠ 0 ∧ b ≠ @[simp] lemma mul_mod_mod (a b c : ℕ) : (a * (b % c)) % c = a * b % c := by rw [mul_mod, mod_mod, ← mul_mod] -@[simp] lemma mod_mul_mod (a b c : ℕ) : (a % c * b) % c = a * b % c := by - rw [mul_mod, mod_mod, ← mul_mod] - lemma pow_mod (a b n : ℕ) : a ^ b % n = (a % n) ^ b % n := by induction b with | zero => rfl diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index f7955d0768276..23cd40d657cba 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -570,7 +570,7 @@ theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits {p : ℕ} (n : ℕ) : theorem digits_two_eq_bits (n : ℕ) : digits 2 n = n.bits.map fun b => cond b 1 0 := by induction' n using Nat.binaryRecFromOne with b n h ih · simp - · rfl + · simp rw [bits_append_bit _ _ fun hn => absurd hn h] cases b · rw [digits_def' one_lt_two] diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index a00644b704354..1285453521a3b 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -13,7 +13,7 @@ import Mathlib.Tactic.IntervalCases # Basic lemmas on prime factorizations -/ -open Nat Finset List Finsupp +open Finset List Finsupp namespace Nat variable {a b m n p : ℕ} @@ -455,6 +455,10 @@ theorem setOf_pow_dvd_eq_Icc_factorization {n p : ℕ} (pp : p.Prime) (hn : n ext simp [Nat.lt_succ_iff, one_le_iff_ne_zero, pp.pow_dvd_iff_le_factorization hn] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The set of positive powers of prime `p` that divide `n` is exactly the set of positive natural numbers up to `n.factorization p`. -/ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : @@ -462,7 +466,7 @@ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : rcases eq_or_ne n 0 with (rfl | hn) · simp ext x - simp only [mem_Icc, Finset.mem_filter, mem_Ico, and_assoc, and_congr_right_iff, + simp only [mem_Icc, Finset.mem_filter, mem_Ico, _root_.and_assoc, and_congr_right_iff, pp.pow_dvd_iff_le_factorization hn, iff_and_self] exact fun _ H => lt_of_le_of_lt H (factorization_lt p hn) diff --git a/Mathlib/Data/Nat/Factorization/Defs.lean b/Mathlib/Data/Nat/Factorization/Defs.lean index 4ff0ac2cf90f5..fb64f3dd9d2d6 100644 --- a/Mathlib/Data/Nat/Factorization/Defs.lean +++ b/Mathlib/Data/Nat/Factorization/Defs.lean @@ -86,7 +86,7 @@ alias factorization_eq_factors_multiset := factorization_eq_primeFactorsList_mul theorem Prime.factorization_pos_of_dvd {n p : ℕ} (hp : p.Prime) (hn : n ≠ 0) (h : p ∣ n) : 0 < n.factorization p := by - rwa [← primeFactorsList_count_eq, count_pos_iff_mem, mem_primeFactorsList_iff_dvd hn hp] + rwa [← primeFactorsList_count_eq, count_pos_iff, mem_primeFactorsList_iff_dvd hn hp] theorem multiplicity_eq_factorization {n p : ℕ} (pp : p.Prime) (hn : n ≠ 0) : multiplicity p n = n.factorization p := by diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index eb2ae3bc58f52..744b67ca07c10 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -121,11 +121,15 @@ lemma log_le_self (b x : ℕ) : log b x ≤ x := theorem lt_pow_succ_log_self {b : ℕ} (hb : 1 < b) (x : ℕ) : x < b ^ (log b x).succ := lt_pow_of_log_lt hb (lt_succ_self _) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : log b n = m ↔ b ^ m ≤ n ∧ n < b ^ (m + 1) := by rcases em (1 < b ∧ n ≠ 0) with (⟨hb, hn⟩ | hbn) - · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, and_comm] <;> - assumption + · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, + _root_.and_comm] <;> assumption have hm : m ≠ 0 := h.resolve_right hbn rw [not_and_or, not_lt, Ne, not_not] at hbn rcases hbn with (hb | rfl) diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index 125eeae18ef2b..418d1347bbc66 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -30,9 +30,13 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b · simp [WithBot.add_bot] simp [← WithBot.coe_add, add_eq_zero_iff_of_nonneg] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by cases n - · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, or_self] + · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, _root_.or_self] cases m · simp [WithBot.add_bot] simp [← WithBot.coe_add, Nat.add_eq_one_iff] diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index ac946c8f455bc..66c8cae9f974a 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -867,7 +867,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by · rfl · rw [PosNum.cast_bit1, ← two_mul, ← congr_fun Nat.bit_true, Nat.testBit_bit_zero] · rw [PosNum.cast_bit0, ← two_mul, ← congr_fun Nat.bit_false, Nat.testBit_bit_zero] - · simp + · simp [Nat.testBit_add_one] · rw [PosNum.cast_bit1, ← two_mul, ← congr_fun Nat.bit_true, Nat.testBit_bit_succ, IH] · rw [PosNum.cast_bit0, ← two_mul, ← congr_fun Nat.bit_false, Nat.testBit_bit_succ, IH] diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index 138ba4471da95..b0bab527487d9 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -137,29 +137,12 @@ variable {p : α → Prop} (f : ∀ a : α, p a → β) (x : Option α) theorem pbind_eq_bind (f : α → Option β) (x : Option α) : (x.pbind fun a _ ↦ f a) = x.bind f := by cases x <;> simp only [pbind, none_bind', some_bind'] -theorem map_bind {α β γ} (f : β → γ) (x : Option α) (g : α → Option β) : - Option.map f (x >>= g) = x >>= fun a ↦ Option.map f (g a) := by - simp only [← map_eq_map, ← bind_pure_comp, LawfulMonad.bind_assoc] - theorem map_bind' (f : β → γ) (x : Option α) (g : α → Option β) : Option.map f (x.bind g) = x.bind fun a ↦ Option.map f (g a) := by cases x <;> simp -theorem map_pbind (f : β → γ) (x : Option α) (g : ∀ a, a ∈ x → Option β) : - Option.map f (x.pbind g) = x.pbind fun a H ↦ Option.map f (g a H) := by - cases x <;> simp only [pbind, map_none'] - theorem pbind_map (f : α → β) (x : Option α) (g : ∀ b : β, b ∈ x.map f → Option γ) : pbind (Option.map f x) g = x.pbind fun a h ↦ g (f a) (mem_map_of_mem _ h) := by cases x <;> rfl -@[simp] -theorem pmap_none (f : ∀ a : α, p a → β) {H} : pmap f (@none α) H = none := - rfl - -@[simp] -theorem pmap_some (f : ∀ a : α, p a → β) {x : α} (h : p x) : - pmap f (some x) = fun _ ↦ some (f x h) := - rfl - theorem mem_pmem {a : α} (h : ∀ a ∈ x, p a) (ha : a ∈ x) : f a (h a ha) ∈ pmap f x h := by rw [mem_def] at ha ⊢ subst ha @@ -208,24 +191,6 @@ theorem pbind_eq_some {f : ∀ a : α, a ∈ x → Option β} {y : β} : simp only [mem_def, Option.some_inj] at H simpa [H] using hz --- Porting note: Can't simp tag this anymore because `pmap` simplifies --- @[simp] -theorem pmap_eq_none_iff {h} : pmap f x h = none ↔ x = none := by cases x <;> simp - --- Porting note: Can't simp tag this anymore because `pmap` simplifies --- @[simp] -theorem pmap_eq_some_iff {hf} {y : β} : - pmap f x hf = some y ↔ ∃ (a : α) (H : x = some a), f a (hf a H) = y := by - rcases x with (_|x) - · simp only [not_mem_none, exists_false, pmap, not_false_iff, exists_prop_of_false, reduceCtorEq] - · constructor - · intro h - simp only [pmap, Option.some_inj] at h - exact ⟨x, rfl, h⟩ - · rintro ⟨a, H, rfl⟩ - simp only [mem_def, Option.some_inj] at H - simp only [H, pmap] - -- Porting note: Can't simp tag this anymore because `join` and `pmap` simplify -- @[simp] theorem join_pmap_eq_pmap_join {f : ∀ a, p a → β} {x : Option (Option α)} (H) : diff --git a/Mathlib/Data/Option/Defs.lean b/Mathlib/Data/Option/Defs.lean index e140c43e996c9..c15bedf853135 100644 --- a/Mathlib/Data/Option/Defs.lean +++ b/Mathlib/Data/Option/Defs.lean @@ -75,10 +75,6 @@ abbrev iget [Inhabited α] : Option α → α theorem iget_some [Inhabited α] {a : α} : (some a).iget = a := rfl -@[simp] -theorem mem_toList {a : α} {o : Option α} : a ∈ toList o ↔ a ∈ o := by - cases o <;> simp [toList, eq_comm] - instance liftOrGet_isCommutative (f : α → α → α) [Std.Commutative f] : Std.Commutative (liftOrGet f) := ⟨fun a b ↦ by cases a <;> cases b <;> simp [liftOrGet, Std.Commutative.comm]⟩ diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 48a139e2ee364..06b25c58fc953 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -378,7 +378,7 @@ theorem Sized.rotateR_size {l x r} (hl : Sized l) : rw [← size_dual, dual_rotateR, hl.dual.rotateL_size, size_dual, size_dual, add_comm (size l)] theorem Sized.balance' {l x r} (hl : @Sized α l) (hr : Sized r) : Sized (balance' l x r) := by - unfold balance'; split_ifs + unfold Ordnode.balance'; split_ifs · exact hl.node' hr · exact hl.rotateL hr · exact hl.rotateR hr @@ -1258,7 +1258,7 @@ theorem Valid'.glue_aux {l r o₁ o₂} (hl : Valid' o₁ l o₂) (hr : Valid' o suffices H : _ by refine ⟨Valid'.balanceL (hl.of_lt ?_ ?_) v H, ?_⟩ · refine @findMin'_all (P := fun a : α => Bounded nil o₁ (a : WithBot α)) - rl rx (sep.2.1.1.imp ?_) hr.1.1.to_nil + _ rl rx (sep.2.1.1.imp ?_) hr.1.1.to_nil exact fun y h => hl.1.1.to_nil.mono_right (le_of_lt h) · exact @findMin'_all _ (fun a => All (· < a) (.node ls ll lx lr)) rl rx diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 73fe251a4dda5..699b8b76a816f 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -470,6 +470,10 @@ theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (F.map dsimp only [PFunctor.map] congr +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx : Agree' n z x) (hy : Agree' n z y) (hrec : ∀ ps : Path F, n = ps.length → iselect ps x = iselect ps y) : x.approx (n + 1) = y.approx (n + 1) := by diff --git a/Mathlib/Data/PNat/Defs.lean b/Mathlib/Data/PNat/Defs.lean index 3004d33fc9df7..6eb644184ac56 100644 --- a/Mathlib/Data/PNat/Defs.lean +++ b/Mathlib/Data/PNat/Defs.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Neil Strickland -/ -import Mathlib.Algebra.NeZero import Mathlib.Data.Nat.Defs import Mathlib.Order.Basic import Mathlib.Order.TypeTags diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index f30532075537d..125f987461805 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -186,12 +186,12 @@ theorem snd_eq_iff : ∀ {p : α × β} {x : β}, p.2 = x ↔ p = (p.1, x) variable {r : α → α → Prop} {s : β → β → Prop} {x y : α × β} -lemma lex_iff : Prod.Lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def _ _ +lemma lex_iff : Prod.Lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def instance Lex.decidable [DecidableEq α] (r : α → α → Prop) (s : β → β → Prop) [DecidableRel r] [DecidableRel s] : DecidableRel (Prod.Lex r s) := - fun _ _ ↦ decidable_of_decidable_of_iff (lex_def r s).symm + fun _ _ ↦ decidable_of_decidable_of_iff lex_def.symm @[refl] theorem Lex.refl_left (r : α → α → Prop) (s : β → β → Prop) [IsRefl α r] : ∀ x, Prod.Lex r s x x diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index 1a8b1f220c0f8..f8116316921ee 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -48,11 +48,11 @@ instance instLT (α β : Type*) [LT α] [LT β] : LT (α ×ₗ β) where lt := P theorem le_iff [LT α] [LE β] (a b : α × β) : toLex a ≤ toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 ≤ b.2 := - Prod.lex_def (· < ·) (· ≤ ·) + Prod.lex_def theorem lt_iff [LT α] [LT β] (a b : α × β) : toLex a < toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 < b.2 := - Prod.lex_def (· < ·) (· < ·) + Prod.lex_def example (x : α) (y : β) : toLex (x, y) = toLex (x, y) := rfl diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index f684d39ce96de..0aa562b90bce1 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -53,7 +53,7 @@ theorem num_mk (n d : ℤ) : (n /. d).num = d.sign * n / n.gcd d := by have (m : ℕ) : Int.natAbs (m + 1) = m + 1 := by rw [← Nat.cast_one, ← Nat.cast_add, Int.natAbs_cast] rcases d with ((_ | _) | _) <;> - rw [← Int.div_eq_ediv_of_dvd] <;> + rw [← Int.tdiv_eq_ediv_of_dvd] <;> simp [divInt, mkRat, Rat.normalize, Nat.succPNat, Int.sign, Int.gcd, Int.zero_ediv, Int.ofNat_dvd_left, Nat.gcd_dvd_left, this] @@ -195,7 +195,7 @@ theorem div_int_inj {a b c d : ℤ} (hb0 : 0 < b) (hd0 : 0 < d) (h1 : Nat.Coprim theorem intCast_div_self (n : ℤ) : ((n / n : ℤ) : ℚ) = n / n := by by_cases hn : n = 0 · subst hn - simp only [Int.cast_zero, Int.zero_div, zero_div, Int.ediv_zero] + simp only [Int.cast_zero, Int.zero_tdiv, zero_div, Int.ediv_zero] · have : (n : ℚ) ≠ 0 := by rwa [← coe_int_inj] at hn simp only [Int.ediv_self hn, Int.cast_one, Ne, not_false_iff, div_self this] @@ -247,9 +247,9 @@ theorem inv_intCast_num (a : ℤ) : (a : ℚ)⁻¹.num = Int.sign a := by rcases lt_trichotomy a 0 with lt | rfl | gt · obtain ⟨a, rfl⟩ : ∃ b, -b = a := ⟨-a, a.neg_neg⟩ simp at lt - simp [Rat.inv_neg, inv_intCast_num_of_pos lt, (Int.sign_eq_one_iff_pos _).mpr lt] - · rfl - · simp [inv_intCast_num_of_pos gt, (Int.sign_eq_one_iff_pos _).mpr gt] + simp [Rat.inv_neg, inv_intCast_num_of_pos lt, Int.sign_eq_one_iff_pos.mpr lt] + · simp + · simp [inv_intCast_num_of_pos gt, Int.sign_eq_one_iff_pos.mpr gt] @[simp] theorem inv_natCast_num (a : ℕ) : (a : ℚ)⁻¹.num = Int.sign a := @@ -268,7 +268,7 @@ theorem inv_intCast_den (a : ℤ) : (a : ℚ)⁻¹.den = if a = 0 then 1 else a. rw [if_neg (by omega)] simp only [Int.cast_neg, Rat.inv_neg, neg_den, inv_intCast_den_of_pos lt, Int.natAbs_neg] exact Int.eq_natAbs_of_zero_le (by omega) - · rfl + · simp · rw [if_neg (by omega)] simp only [inv_intCast_den_of_pos gt] exact Int.eq_natAbs_of_zero_le (by omega) diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 8a167a30145c9..4817d9024ec51 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -717,12 +717,9 @@ theorem head_terminates_of_head_tail_terminates (s : WSeq α) [T : Terminates (h Terminates (head s) := (head_terminates_iff _).2 <| by rcases (head_terminates_iff _).1 T with ⟨⟨a, h⟩⟩ - simp? [tail] at h says simp only [tail, destruct_flatten] at h + simp? [tail] at h says simp only [tail, destruct_flatten, bind_map_left] at h rcases exists_of_mem_bind h with ⟨s', h1, _⟩ - unfold Functor.map at h1 - exact - let ⟨t, h3, _⟩ := Computation.exists_of_mem_map h1 - Computation.terminates_of_mem h3 + exact terminates_of_mem h1 theorem destruct_some_of_destruct_tail_some {s : WSeq α} {a} (h : some a ∈ destruct (tail s)) : ∃ a', some a' ∈ destruct s := by diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 25b578031b0f6..5a13d62ad7baf 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -118,7 +118,6 @@ lemma memPartitionSet_succ (f : ℕ → Set α) (n : ℕ) (a : α) [Decidable (a memPartitionSet f (n + 1) a = if a ∈ f n then memPartitionSet f n a ∩ f n else memPartitionSet f n a \ f n := by simp [memPartitionSet] - congr lemma memPartitionSet_mem (f : ℕ → Set α) (n : ℕ) (a : α) : memPartitionSet f n a ∈ memPartition f n := by diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index ca1d0ac665068..4f812be71f3cd 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -420,7 +420,7 @@ def sigmaQuotientEquivOfLe {r s : Setoid α} (hle : r ≤ s) : (Σ q : Quotient s, Quotient (r.comap (Subtype.val : Quotient.mk s ⁻¹' {q} → α))) ≃ Quotient r := .trans (.symm <| .sigmaCongrRight fun _ ↦ .subtypeQuotientEquivQuotientSubtype - (s₁ := r) (s₂ := r.comap Subtype.val) _ (fun _ ↦ Iff.rfl) fun _ _ ↦ Iff.rfl) + (s₁ := r) (s₂ := r.comap Subtype.val) _ _ (fun _ ↦ Iff.rfl) fun _ _ ↦ Iff.rfl) (.sigmaFiberEquiv fun a ↦ a.lift (Quotient.mk s) fun _ _ h ↦ Quotient.sound <| hle h) end Setoid diff --git a/Mathlib/Data/Vector/Defs.lean b/Mathlib/Data/Vector/Defs.lean index 81ec5ceb94570..8c08c0254b502 100644 --- a/Mathlib/Data/Vector/Defs.lean +++ b/Mathlib/Data/Vector/Defs.lean @@ -119,7 +119,7 @@ def take (i : ℕ) : Vector α n → Vector α (min i n) /-- Remove the element at position `i` from a vector of length `n`. -/ def eraseIdx (i : Fin n) : Vector α n → Vector α (n - 1) - | ⟨l, p⟩ => ⟨List.eraseIdx l i.1, by rw [l.length_eraseIdx] <;> rw [p]; exact i.2⟩ + | ⟨l, p⟩ => ⟨List.eraseIdx l i.1, by rw [l.length_eraseIdx_of_lt] <;> rw [p]; exact i.2⟩ @[deprecated (since := "2024-05-04")] alias removeNth := eraseIdx diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 6977197f083cc..76e27012548a5 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -104,7 +104,7 @@ namespace ZMod instance instUnique : Unique (ZMod 1) := Fin.uniqueFinOne instance fintype : ∀ (n : ℕ) [NeZero n], Fintype (ZMod n) - | 0, h => (h.ne rfl).elim + | 0, h => (h.ne _ rfl).elim | n + 1, _ => Fin.fintype (n + 1) instance infinite : Infinite (ZMod 0) := diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 3f47192c4cc47..160a6606f5e0f 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -338,7 +338,7 @@ theorem not_isPeriodicPt_of_pos_of_lt_minimalPeriod : theorem IsPeriodicPt.minimalPeriod_dvd (hx : IsPeriodicPt f n x) : minimalPeriod f x ∣ n := (eq_or_lt_of_le <| n.zero_le).elim (fun hn0 => hn0 ▸ dvd_zero _) fun hn0 => -- Porting note: `Nat.dvd_iff_mod_eq_zero` gained explicit arguments - (Nat.dvd_iff_mod_eq_zero _ _).2 <| + Nat.dvd_iff_mod_eq_zero.2 <| (hx.mod <| isPeriodicPt_minimalPeriod f x).eq_zero_of_lt_minimalPeriod <| Nat.mod_lt _ <| hx.minimalPeriod_pos hn0 @@ -433,7 +433,7 @@ theorem periodicOrbit_length : (periodicOrbit f x).length = minimalPeriod f x := @[simp] theorem periodicOrbit_eq_nil_iff_not_periodic_pt : periodicOrbit f x = Cycle.nil ↔ x ∉ periodicPts f := by - simp only [periodicOrbit.eq_1, Cycle.coe_eq_nil, List.map_eq_nil, List.range_eq_nil] + simp only [periodicOrbit.eq_1, Cycle.coe_eq_nil, List.map_eq_nil_iff, List.range_eq_nil] exact minimalPeriod_eq_zero_iff_nmem_periodicPts theorem periodicOrbit_eq_nil_of_not_periodic_pt (h : x ∉ periodicPts f) : diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 21d9418e249a5..787429c5e3827 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -79,7 +79,7 @@ theorem spanEval_ne_top : spanEval k ≠ ⊤ := by rw [map_one, Finsupp.linearCombination_apply, Finsupp.sum, map_sum, Finset.sum_eq_zero] at hv · exact zero_ne_one hv intro j hj - rw [smul_eq_mul, map_mul, toSplittingField_evalXSelf (s := v.support) hj, + rw [smul_eq_mul, map_mul, toSplittingField_evalXSelf _ (s := v.support) hj, mul_zero] /-- A random maximal ideal that contains `spanEval k` -/ diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 939bc161ec3f8..671a4b1bb374a 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -272,6 +272,11 @@ theorem Polynomial.separable_X_pow_sub_C_of_irreducible : (X ^ n - C a).Separabl AdjoinRoot.algebraMap_eq, X_pow_sub_C_eq_prod (hζ.map_of_injective (algebraMap K _).injective) hn (root_X_pow_sub_C_pow n a), separable_prod_X_sub_C_iff'] + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 we need to provide this helper instance. + -/ + have : MonoidHomClass (K →+* K[n√a]) K K[n√a] := inferInstance exact (hζ.map_of_injective (algebraMap K K[n√a]).injective).injOn_pow_mul (root_X_pow_sub_C_ne_zero (lt_of_le_of_ne (show 1 ≤ n from hn) (Ne.symm hn')) _) diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index c880a81782d7f..93c3226de53d7 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -166,7 +166,7 @@ variable {F K} theorem isPurelyInseparable_iff : IsPurelyInseparable F E ↔ ∀ x : E, IsIntegral F x ∧ (IsSeparable F x → x ∈ (algebraMap F E).range) := - ⟨fun h x ↦ ⟨h.isIntegral' x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ + ⟨fun h x ↦ ⟨h.isIntegral' _ x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ /-- Transfer `IsPurelyInseparable` across an `AlgEquiv`. -/ theorem AlgEquiv.isPurelyInseparable (e : K ≃ₐ[F] E) [IsPurelyInseparable F K] : diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index c6a87d0dbf1fb..fd852c7cdf7d7 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -652,7 +652,7 @@ theorem IsSeparable.tower_bot {x : K} (h : IsSeparable F (algebraMap K E x)) : I variable (K E) in theorem Algebra.isSeparable_tower_bot_of_isSeparable [h : Algebra.IsSeparable F E] : Algebra.IsSeparable F K := - ⟨fun _ ↦ IsSeparable.tower_bot (h.isSeparable _)⟩ + ⟨fun _ ↦ IsSeparable.tower_bot (h.isSeparable _ _)⟩ end IsScalarTower diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 6ebcf74215644..fb1ec8ab2843c 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -208,7 +208,7 @@ def embEquivOfAdjoinSplits {S : Set E} (hS : adjoin F S = ⊤) (hS ▸ isAlgebraic_adjoin (S := S) fun x hx ↦ (hK x hx).1) have halg := (topEquiv (F := F) (E := E)).isAlgebraic Classical.choice <| Function.Embedding.antisymm - (halg.algHomEmbeddingOfSplits (fun _ ↦ splits_of_mem_adjoin F (S := S) hK (hS ▸ mem_top)) _) + (halg.algHomEmbeddingOfSplits (fun _ ↦ splits_of_mem_adjoin F E (S := S) hK (hS ▸ mem_top)) _) (halg.algHomEmbeddingOfSplits (fun _ ↦ IsAlgClosed.splits_codomain _) _) /-- The `Field.finSepDegree F E` is equal to the cardinality of `E →ₐ[F] K` diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 5c5349f4adc95..a966c23b3b918 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -167,7 +167,7 @@ noncomputable def invApp (U : Opens X) : @[simp, reassoc] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = + X.presheaf.map i ≫ H.invApp _ (unop V) = invApp f (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := by simp only [invApp, ← Category.assoc] rw [IsIso.comp_inv_eq] @@ -179,11 +179,11 @@ theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : instance (U : Opens X) : IsIso (invApp f U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (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])) := by - rw [← cancel_epi (H.invApp U), IsIso.hom_inv_id] + rw [← cancel_epi (H.invApp _ U), IsIso.hom_inv_id] delta invApp simp [← Functor.map_comp] @@ -195,7 +195,7 @@ theorem invApp_app (U : Opens X) : @[simp, reassoc] theorem app_invApp (U : Opens Y) : - f.c.app (op U) ≫ H.invApp ((Opens.map f.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.base U.1)).op : op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.base).obj U))) := by @@ -244,7 +244,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) : - (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := by + (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := by delta invApp rw [IsIso.comp_inv_eq, Category.id_comp] change X.presheaf.map _ = X.presheaf.map _ @@ -290,7 +290,7 @@ def pullbackConeOfLeftFst : base := pullback.fst _ _ c := { app := fun U => - hf.invApp (unop U) ≫ + hf.invApp _ (unop U) ≫ g.c.app (op (hf.base_open.isOpenMap.functor.obj (unop U))) ≫ Y.presheaf.map (eqToHom @@ -772,21 +772,21 @@ noncomputable def invApp (U : Opens X) : @[reassoc (attr := simp)] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = - H.invApp (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := + X.presheaf.map i ≫ H.invApp _ (unop V) = + H.invApp _ (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := PresheafedSpace.IsOpenImmersion.inv_naturality f i -instance (U : Opens X) : IsIso (H.invApp U) := by delta invApp; infer_instance +instance (U : Opens X) : IsIso (H.invApp _ U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (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.inv_invApp f U @[reassoc (attr := simp)] theorem invApp_app (U : Opens X) : - H.invApp U ≫ f.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 U @@ -794,7 +794,7 @@ attribute [elementwise] invApp_app @[reassoc (attr := simp)] theorem app_invApp (U : Opens Y) : - f.c.app (op U) ≫ H.invApp ((Opens.map f.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.base U.1)).op : op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.base).obj U))) := @@ -818,7 +818,7 @@ instance ofRestrict {X : TopCat} (Y : SheafedSpace C) {f : X ⟶ Y.carrier} @[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) : - (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := + (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U /-- An open immersion is an iso if the underlying continuous map is epi. -/ @@ -1154,7 +1154,7 @@ 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) [stalk_iso : ∀ x : X.1, IsIso (f.stalkMap x)] : LocallyRingedSpace.IsOpenImmersion f := - SheafedSpace.IsOpenImmersion.of_stalk_iso hf (H := stalk_iso) + SheafedSpace.IsOpenImmersion.of_stalk_iso _ hf (H := stalk_iso) end OfStalkIso @@ -1180,21 +1180,21 @@ noncomputable def invApp (U : Opens X) : @[reassoc (attr := simp)] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = - H.invApp (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := + X.presheaf.map i ≫ H.invApp _ (unop V) = + H.invApp _ (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := PresheafedSpace.IsOpenImmersion.inv_naturality f.1 i -instance (U : Opens X) : IsIso (H.invApp U) := by delta invApp; infer_instance +instance (U : Opens X) : IsIso (H.invApp _ U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (H.invApp _ U) = f.1.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.1.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 @@ -1202,7 +1202,7 @@ 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.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.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))) := @@ -1211,7 +1211,7 @@ theorem app_invApp (U : Opens Y) : /-- 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) = + f.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.base).obj U) = Y.presheaf.map (eqToHom <| le_antisymm (Set.image_preimage_subset f.1.base U.1) <| @@ -1226,7 +1226,7 @@ instance ofRestrict {X : TopCat} (Y : LocallyRingedSpace) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp (X : LocallyRingedSpace) {Y : TopCat} {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : - (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := + (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U instance stalk_iso (x : X) : IsIso (f.stalkMap x) := diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index b5ac976467db8..4d9076090a048 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -149,7 +149,7 @@ theorem pullback_base (i j k : D.J) (S : Set (D.V (i, j)).carrier) : /-- The red and the blue arrows in ![this diagram](https://i.imgur.com/0GiBUh6.png) commute. -/ @[simp, reassoc] theorem f_invApp_f_app (i j k : D.J) (U : Opens (D.V (i, j)).carrier) : - (D.f_open i j).invApp U ≫ (D.f i k).c.app _ = + (D.f_open i j).invApp _ U ≫ (D.f i k).c.app _ = (π₁ i, j, k).c.app (op U) ≫ (π₂⁻¹ i, j, k) (unop _) ≫ (D.V _).presheaf.map @@ -161,7 +161,7 @@ theorem f_invApp_f_app (i j k : D.J) (U : Opens (D.V (i, j)).carrier) : apply pullback_base)) := by have := PresheafedSpace.congr_app (@pullback.condition _ _ _ _ _ (D.f i j) (D.f i k) _) dsimp only [comp_c_app] at this - rw [← cancel_epi (inv ((D.f_open i j).invApp U)), IsIso.inv_hom_id_assoc, + rw [← cancel_epi (inv ((D.f_open i j).invApp _ U)), IsIso.inv_hom_id_assoc, IsOpenImmersion.inv_invApp] simp_rw [Category.assoc] erw [(π₁ i, j, k).c.naturality_assoc, reassoc_of% this, ← Functor.map_comp_assoc, @@ -271,7 +271,7 @@ def opensImagePreimageMap (i j : D.J) (U : Opens (D.U i).carrier) : (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding 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 _) ≫ + (D.f_open j i).invApp _ (unop _) ≫ (𝖣.U j).presheaf.map (eqToHom (D.ι_image_preimage_eq i j U)).op theorem opensImagePreimageMap_app' (i j k : D.J) (U : Opens (D.U i).carrier) : @@ -436,19 +436,9 @@ abbrev ιInvAppπEqMap {i : D.J} (U : Opens (D.U i).carrier) := theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : D.diagramOverOpenπ U i ≫ D.ιInvAppπEqMap U ≫ D.ιInvApp U ≫ D.diagramOverOpenπ U j = D.diagramOverOpenπ U j := by - -- Porting note: originally, the proof of monotonicity was left a blank and proved in the end - -- but Lean 4 doesn't like this any more, so the proof is restructured - rw [← @cancel_mono (f := (componentwiseDiagram 𝖣.diagram.multispan _).map - (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) _ _ (by - rw [Category.comp_id] - apply (config := { allowSynthFailures := true }) mono_comp - change Mono ((_ ≫ D.f j i).c.app _) - rw [comp_c_app] - apply (config := { allowSynthFailures := true }) mono_comp - · erw [D.ι_image_preimage_eq i j U] - infer_instance - · have : IsIso (D.t i j).c := by apply c_isIso_of_iso - infer_instance)] + rw [← @cancel_mono + (f := (componentwiseDiagram 𝖣.diagram.multispan _).map + (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) ..] simp_rw [Category.assoc] rw [limit.w_assoc] erw [limit.lift_π_assoc] @@ -466,6 +456,15 @@ theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : convert limit.w (componentwiseDiagram 𝖣.diagram.multispan _) (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) + · rw [Category.comp_id] + apply (config := { allowSynthFailures := true }) mono_comp + change Mono ((_ ≫ D.f j i).c.app _) + rw [comp_c_app] + apply (config := { allowSynthFailures := true }) mono_comp + · erw [D.ι_image_preimage_eq i j U] + infer_instance + · have : IsIso (D.t i j).c := by apply c_isIso_of_iso + infer_instance /-- `ιInvApp` is the inverse of `D.ι i` on `U`. -/ theorem π_ιInvApp_eq_id (i : D.J) (U : Opens (D.U i).carrier) : diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index c02743ac71a50..40bcb398e5f67 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -131,6 +131,7 @@ lemma commProb_odd {n : ℕ} (hn : Odd n) : qify [show 2 ∣ n + 3 by rw [Nat.dvd_iff_mod_eq_zero, Nat.add_mod, Nat.odd_iff.mp hn]] rw [div_div, ← mul_assoc] congr + norm_num private lemma div_two_lt {n : ℕ} (h0 : n ≠ 0) : n / 2 < n := Nat.div_lt_self (Nat.pos_of_ne_zero h0) (lt_add_one 1) diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index e1096e6829211..4dfea2fbf13f6 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -128,7 +128,7 @@ theorem not_step_nil : ¬Step [] L := by generalize h' : [] = L' intro h cases' h with L₁ L₂ - simp [List.nil_eq_append] at h' + simp [List.nil_eq_append_iff] at h' @[to_additive] theorem Step.cons_left_iff {a : α} {b : Bool} : @@ -287,7 +287,8 @@ theorem red_iff_irreducible {x1 b1 x2 b2} (h : (x1, b1) ≠ (x2, b2)) : generalize eq : [(x1, not b1), (x2, b2)] = L' intro L h' cases h' - simp [List.cons_eq_append, List.nil_eq_append] at eq + simp only [List.cons_eq_append_iff, List.cons.injEq, Prod.mk.injEq, and_false, + List.nil_eq_append_iff, exists_const, or_self, or_false, List.cons_ne_nil] at eq rcases eq with ⟨rfl, ⟨rfl, rfl⟩, ⟨rfl, rfl⟩, rfl⟩ simp at h diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index f036e80ab2244..7bf00ac53f9d3 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -404,6 +404,12 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] convert one_mem H rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), mem_stabilizer_iff, mul_smul, (exists_smul_eq α (g⁻¹ • x) x).choose_spec] + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 we need to search for this instance explicitly. + TODO: change `convert` to more agressively solve such goals with `infer_instance` itself. + -/ + infer_instance end MulAction diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 87db896052260..9e33d48b21996 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -717,7 +717,7 @@ automatic in the case of a finite cancellative monoid. -/ `addOrderOf_nsmul` but with one assumption less which is automatic in the case of a finite cancellative additive monoid."] theorem orderOf_pow (x : G) : orderOf (x ^ n) = orderOf x / gcd (orderOf x) n := - (isOfFinOrder_of_finite _).orderOf_pow _ + (isOfFinOrder_of_finite _).orderOf_pow .. @[to_additive] theorem mem_powers_iff_mem_range_orderOf [DecidableEq G] : diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 57afc2d245faf..c16d73ee548bd 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -316,7 +316,7 @@ theorem toList_formPerm_nontrivial (l : List α) (hl : 2 ≤ l.length) (hn : Nod · refine ext_getElem (by simp) fun k hk hk' => ?_ simp only [get_eq_getElem, formPerm_pow_apply_getElem _ hn, zero_add, getElem_map, getElem_range, Nat.mod_eq_of_lt hk'] - · simpa [hs] using get_mem _ _ _ + · simp [hs] theorem toList_formPerm_isRotated_self (l : List α) (hl : 2 ≤ l.length) (hn : Nodup l) (x : α) (hx : x ∈ l) : toList (formPerm l) x ~r l := by diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index a7d22e67b1e55..269920ef0aab3 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -417,7 +417,7 @@ theorem sign_trans_trans_symm [DecidableEq β] [Fintype β] (f : Perm β) (e : theorem sign_prod_list_swap {l : List (Perm α)} (hl : ∀ g ∈ l, IsSwap g) : sign l.prod = (-1) ^ l.length := by have h₁ : l.map sign = List.replicate l.length (-1) := - List.eq_replicate.2 + List.eq_replicate_iff.2 ⟨by simp, fun u hu => let ⟨g, hg⟩ := List.mem_map.1 hu hg.2 ▸ (hl _ hg.1).sign_eq⟩ diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index 7ae6931a6aa29..9f90f39e9ee1f 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -226,7 +226,7 @@ theorem MultilinearMap.domCoprod_alternization [DecidableEq ιa] [DecidableEq ι -- unfold the quotient mess left by `Finset.sum_partition` -- Porting note: Was `conv in .. => ..`. erw - [@Finset.filter_congr _ _ (fun a => @Quotient.decidableEq _ _ + [@Finset.filter_congr _ _ _ (fun a => @Quotient.decidableEq _ _ (QuotientGroup.leftRelDecidable (MonoidHom.range (Perm.sumCongrHom ιa ιb))) (Quotient.mk (QuotientGroup.leftRel (MonoidHom.range (Perm.sumCongrHom ιa ιb))) a) (Quotient.mk'' σ)) _ (s := Finset.univ) diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 4359ba5b328d0..e5f6c2be85383 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -154,7 +154,7 @@ theorem evenOdd_induction (n : ZMod 2) {motive : ∀ x, x ∈ evenOdd Q n → Pr motive (ι Q m₁ * ι Q m₂ * x) (zero_add n ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q n) : motive x hx := by - apply Submodule.iSup_induction' (C := motive) _ (range_ι_pow 0 (Submodule.zero_mem _)) add + apply Submodule.iSup_induction' (C := motive) _ _ (range_ι_pow 0 (Submodule.zero_mem _)) add refine Subtype.rec ?_ simp_rw [ZMod.natCast_eq_iff, add_comm n.val] rintro n' ⟨k, rfl⟩ xv @@ -197,7 +197,7 @@ 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 + refine evenOdd_induction _ _ (motive := motive) (fun rx => ?_) add ι_mul_ι_mul x hx rintro ⟨r, rfl⟩ exact algebraMap r @@ -213,7 +213,7 @@ theorem odd_induction {P : ∀ x, x ∈ evenOdd Q 1 → Prop} P (CliffordAlgebra.ι Q m₁ * CliffordAlgebra.ι Q m₂ * x) (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 + 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 rintro ⟨v, rfl⟩ diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 1d2a2669cab75..2b1fa24c031b5 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -183,7 +183,7 @@ def LinearMap.dualMap (f : M₁ →ₗ[R] M₂) : Dual R M₂ →ₗ[R] Dual R M -- Porting note: with reducible def need to specify some parameters to transpose explicitly Module.Dual.transpose (R := R) f -lemma LinearMap.dualMap_eq_lcomp (f : M₁ →ₗ[R] M₂) : f.dualMap = f.lcomp R := rfl +lemma LinearMap.dualMap_eq_lcomp (f : M₁ →ₗ[R] M₂) : f.dualMap = f.lcomp R R := rfl -- Porting note: with reducible def need to specify some parameters to transpose explicitly theorem LinearMap.dualMap_def (f : M₁ →ₗ[R] M₂) : f.dualMap = Module.Dual.transpose (R := R) f := diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index acd30d6e430a9..0b10bbf352e99 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -671,7 +671,7 @@ noncomputable def divisionRingOfFiniteDimensional (F K : Type*) [Field F] [Ring inv x := letI := Classical.decEq K 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 + 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):) inv_zero := dif_pos rfl diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index b814a9748abac..84ce3839bb755 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -495,8 +495,8 @@ theorem adjugate_adjugate (A : Matrix n n α) (h : Fintype.card n ≠ 1) : let A' := mvPolynomialX n n ℤ suffices adjugate (adjugate A') = det A' ^ (Fintype.card n - 2) • A' by rw [← mvPolynomialX_mapMatrix_aeval ℤ A, ← AlgHom.map_adjugate, ← AlgHom.map_adjugate, this, - ← AlgHom.map_det, ← map_pow (MvPolynomial.aeval _), AlgHom.mapMatrix_apply, - AlgHom.mapMatrix_apply, Matrix.map_smul' _ _ _ (_root_.map_mul _)] + ← AlgHom.map_det, ← map_pow (MvPolynomial.aeval fun p : n × n ↦ A p.1 p.2), + AlgHom.mapMatrix_apply, AlgHom.mapMatrix_apply, Matrix.map_smul' _ _ _ (_root_.map_mul _)] have h_card' : Fintype.card n - 2 + 1 = Fintype.card n - 1 := by simp [h_card] have is_reg : IsSMulRegular (MvPolynomial (n × n) ℤ) (det A') := fun x y => mul_left_cancel₀ (det_mvPolynomialX_ne_zero n ℤ) diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index cc04b3b64ac37..4a0d9f53997f4 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -204,7 +204,7 @@ lemma derivative_det_one_add_X_smul_aux {n} (M : Matrix (Fin n) (Fin n) R) : rw [det_eq_zero_of_column_eq_zero 0, eval_zero, mul_zero] intro j rw [submatrix_apply, Fin.succAbove_of_castSucc_lt, one_apply_ne] - · exact (bne_iff_ne (Fin.succ j) (Fin.castSucc 0)).mp rfl + · exact (bne_iff_ne (a := Fin.succ j) (b := Fin.castSucc 0)).mp rfl · rw [Fin.castSucc_zero]; exact lt_of_le_of_ne (Fin.zero_le _) hi.symm · exact fun H ↦ (H <| Finset.mem_univ _).elim diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 9490aa658bb95..c33327a8858ae 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -747,7 +747,7 @@ theorem det_fin_one_of (a : R) : det !![a] = a := theorem det_fin_two (A : Matrix (Fin 2) (Fin 2) R) : det A = A 0 0 * A 1 1 - A 0 1 * A 1 0 := by simp only [det_succ_row_zero, det_unique, Fin.default_eq_zero, submatrix_apply, Fin.succ_zero_eq_one, Fin.sum_univ_succ, Fin.val_zero, Fin.zero_succAbove, univ_unique, - Fin.val_succ, Fin.coe_fin_one, Fin.succ_succAbove_zero, sum_singleton] + Fin.val_succ, Fin.val_eq_zero, Fin.succ_succAbove_zero, sum_singleton] ring @[simp] @@ -763,7 +763,7 @@ theorem det_fin_three (A : Matrix (Fin 3) (Fin 3) R) : simp only [det_succ_row_zero, ← Nat.not_even_iff_odd, submatrix_apply, Fin.succ_zero_eq_one, submatrix_submatrix, det_unique, Fin.default_eq_zero, comp_apply, Fin.succ_one_eq_two, Fin.sum_univ_succ, Fin.val_zero, Fin.zero_succAbove, univ_unique, Fin.val_succ, - Fin.coe_fin_one, Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one, even_add_self] + Fin.val_eq_zero, Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one, even_add_self] ring end Matrix diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 0b3d984857b62..3aeb94be63ed6 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -80,7 +80,7 @@ lemma eq_zero_of_isNilpotent_isSemisimple (hn : IsNilpotent f) (hs : f.IsSemisim have ⟨n, h0⟩ := hn rw [← aeval_X (R := R) f]; rw [← aeval_X_pow (R := R) f] at h0 rw [← RingHom.mem_ker, ← AEval.annihilator_eq_ker_aeval (M := M)] at h0 ⊢ - exact hs.annihilator_isRadical ⟨n, h0⟩ + exact hs.annihilator_isRadical _ _ ⟨n, h0⟩ @[simp] lemma isSemisimple_sub_algebraMap_iff {μ : R} : diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index d06c680fb6859..18c2623b09806 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -220,7 +220,7 @@ theorem tmul_coe_mul_one_tmul {j₁ : ι} (a₁ : A) (b₁ : ℬ j₁) (b₂ : B theorem tmul_one_mul_one_tmul (a₁ : A) (b₂ : B) : (a₁ ᵍ⊗ₜ[R] (1 : B) * (1 : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = (a₁ : A) ᵍ⊗ₜ (b₂ : B) := by convert tmul_coe_mul_zero_coe_tmul 𝒜 ℬ - a₁ (@GradedMonoid.GOne.one _ (ℬ ·) _ _) (@GradedMonoid.GOne.one _ (𝒜 ·) _ _) b₂ + a₁ (GradedMonoid.GOne.one (A := (ℬ ·))) (GradedMonoid.GOne.one (A := (𝒜 ·))) b₂ · rw [SetLike.coe_gOne, mul_one] · rw [SetLike.coe_gOne, one_mul] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 8fd66baa31e7d..78b70b4f49308 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -975,7 +975,7 @@ theorem not_beq_of_ne {α : Type*} [BEq α] [LawfulBEq α] {a b : α} (ne : a fun h => ne (eq_of_beq h) theorem beq_eq_decide {α : Type*} [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by - rw [← beq_iff_eq a b] + rw [← beq_iff_eq (a := a) (b := b)] cases a == b <;> simp @[simp] lemma beq_eq_beq {α β : Type*} [BEq α] [LawfulBEq α] [BEq β] [LawfulBEq β] {a₁ a₂ : α} diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index 60412ec201292..c111c09a26737 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -555,9 +555,13 @@ theorem sequence_mono_nat {r : β → β → Prop} {f : α → β} (hf : Directe · exact (Classical.choose_spec (hf p p)).1 · exact (Classical.choose_spec (hf p a)).1 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem rel_sequence {r : β → β → Prop} {f : α → β} (hf : Directed r f) (a : α) : r (f a) (f (hf.sequence f (encode a + 1))) := by - simp only [Directed.sequence, add_eq, add_zero, encodek, and_self] + simp only [Directed.sequence, add_eq, add_zero, encodek, _root_.and_self] exact (Classical.choose_spec (hf _ a)).2 variable [Preorder β] {f : α → β} diff --git a/Mathlib/Logic/Equiv/Array.lean b/Mathlib/Logic/Equiv/Array.lean index cd9c6fbef7d1f..eb98b0edd323b 100644 --- a/Mathlib/Logic/Equiv/Array.lean +++ b/Mathlib/Logic/Equiv/Array.lean @@ -42,7 +42,7 @@ namespace Equiv /-- The natural equivalence between arrays and lists. -/ def arrayEquivList (α : Type*) : Array α ≃ List α := - ⟨Array.data, Array.mk, fun _ => rfl, fun _ => rfl⟩ + ⟨Array.toList, Array.mk, fun _ => rfl, fun _ => rfl⟩ end Equiv diff --git a/Mathlib/Logic/Nonempty.lean b/Mathlib/Logic/Nonempty.lean index 1a52ca02ab2d4..1c8198b9475d3 100644 --- a/Mathlib/Logic/Nonempty.lean +++ b/Mathlib/Logic/Nonempty.lean @@ -99,13 +99,6 @@ theorem Nonempty.elim_to_inhabited {α : Sort*} [h : Nonempty α] {p : Prop} (f p := h.elim <| f ∘ Inhabited.mk -protected instance Prod.instNonempty {α β} [h : Nonempty α] [h2 : Nonempty β] : Nonempty (α × β) := - h.elim fun g ↦ h2.elim fun g2 ↦ ⟨⟨g, g2⟩⟩ - -protected instance Pi.instNonempty {ι : Sort*} {α : ι → Sort*} [∀ i, Nonempty (α i)] : - Nonempty (∀ i, α i) := - ⟨fun _ ↦ Classical.arbitrary _⟩ - theorem Classical.nonempty_pi {ι} {α : ι → Sort*} : Nonempty (∀ i, α i) ↔ ∀ i, Nonempty (α i) := ⟨fun ⟨f⟩ a ↦ ⟨f a⟩, @Pi.instNonempty _ _⟩ diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean index e0029edb7097f..4bd14a8a98d53 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean @@ -568,7 +568,8 @@ theorem Measurable.isLUB_of_mem {ι} [Countable ι] {f : ι → δ → α} {g g' · simp [hb, hg' hb] rw [this] exact Measurable.piecewise hs measurable_const g'_meas - · let f' : ι → δ → α := fun i ↦ s.piecewise (f i) g' + · have : Nonempty ι := ⟨i⟩ + let f' : ι → δ → α := fun i ↦ s.piecewise (f i) g' suffices ∀ b, IsLUB { a | ∃ i, f' i b = a } (g b) from Measurable.isLUB (fun i ↦ Measurable.piecewise hs (hf i) g'_meas) this intro b @@ -576,14 +577,7 @@ theorem Measurable.isLUB_of_mem {ι} [Countable ι] {f : ι → δ → α} {g g' · have A : ∀ i, f' i b = f i b := fun i ↦ by simp [f', hb] simpa [A] using hg b hb · have A : ∀ i, f' i b = g' b := fun i ↦ by simp [f', hb] - have : {a | ∃ (_i : ι), g' b = a} = {g' b} := by - apply Subset.antisymm - · rintro - ⟨_j, rfl⟩ - simp only [mem_singleton_iff] - · rintro - rfl - exact ⟨i, rfl⟩ - simp only [exists_prop'] at this - simp [A, this, hg' hb, isLUB_singleton] + simp [A, hg' hb, isLUB_singleton] theorem AEMeasurable.isLUB {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, AEMeasurable (f i) μ) (hg : ∀ᵐ b ∂μ, IsLUB { a | ∃ i, f i b = a } (g b)) : diff --git a/Mathlib/MeasureTheory/Measure/Content.lean b/Mathlib/MeasureTheory/Measure/Content.lean index 5306eb0281091..b07a09b9ad57e 100644 --- a/Mathlib/MeasureTheory/Measure/Content.lean +++ b/Mathlib/MeasureTheory/Measure/Content.lean @@ -310,7 +310,7 @@ variable [S : MeasurableSpace G] [BorelSpace G] /-- For the outer measure coming from a content, all Borel sets are measurable. -/ theorem borel_le_caratheodory : S ≤ μ.outerMeasure.caratheodory := by - rw [@BorelSpace.measurable_eq G _ _] + rw [BorelSpace.measurable_eq (α := G)] refine MeasurableSpace.generateFrom_le ?_ intro U hU rw [μ.outerMeasure_caratheodory] diff --git a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean index 2a839ffd63d95..c2c75e666e4b2 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean @@ -140,7 +140,7 @@ theorem card_ring : card Language.ring = 5 := by have : Fintype.card Language.ring.Symbols = 5 := rfl simp [Language.card, this] -open Language ring Structure +open Language Structure /-- A Type `R` is a `CompatibleRing` if it is a structure for the language of rings and this structure is the same as the structure already given on `R` by the classes `Add`, `Mul` etc. diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index e19707451b558..baa0ada8d4271 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -247,14 +247,14 @@ theorem nonempty_of_nonempty_constants [h : Nonempty L.Constants] : Nonempty M : /-- `HomClass L F M N` states that `F` is a type of `L`-homomorphisms. You should extend this typeclass when you extend `FirstOrder.Language.Hom`. -/ -class HomClass (L : outParam Language) (F M N : Type*) +class HomClass (L : outParam Language) (F : Type*) (M N : outParam Type*) [FunLike F M N] [L.Structure M] [L.Structure N] : Prop where map_fun : ∀ (φ : F) {n} (f : L.Functions n) (x), φ (funMap f x) = funMap f (φ ∘ x) map_rel : ∀ (φ : F) {n} (r : L.Relations n) (x), RelMap r x → RelMap r (φ ∘ x) /-- `StrongHomClass L F M N` states that `F` is a type of `L`-homomorphisms which preserve relations in both directions. -/ -class StrongHomClass (L : outParam Language) (F M N : Type*) +class StrongHomClass (L : outParam Language) (F : Type*) (M N : outParam Type*) [FunLike F M N] [L.Structure M] [L.Structure N] : Prop where map_fun : ∀ (φ : F) {n} (f : L.Functions n) (x), φ (funMap f x) = funMap f (φ ∘ x) map_rel : ∀ (φ : F) {n} (r : L.Relations n) (x), RelMap r (φ ∘ x) ↔ RelMap r x diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index fab8363491989..a9157aba577e8 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -83,7 +83,7 @@ theorem listDecode_encode_list (l : List (L.Term α)) : simp only [h, length_append, length_map, length_finRange, le_add_iff_nonneg_right, _root_.zero_le, ↓reduceDIte, getElem_fin, cons.injEq, func.injEq, heq_eq_eq, true_and] refine ⟨funext (fun i => ?_), ?_⟩ - · rw [List.getElem_append, List.getElem_map, List.getElem_finRange] + · rw [List.getElem_append_left, List.getElem_map, List.getElem_finRange] simp only [length_map, length_finRange, i.2] · simp only [length_map, length_finRange, drop_left'] @@ -244,7 +244,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : simp only [Option.join, map_append, map_map, Option.bind_eq_some, id, exists_eq_right, get?_eq_some, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ - rw [get_eq_getElem, getElem_append, getElem_map] + rw [get_eq_getElem, getElem_append_left, getElem_map] · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index fb5f3691aab1a..7fa52ce76edb5 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -576,8 +576,8 @@ theorem relabel_sum_inl (φ : L.BoundedFormula α n) : | falsum => rfl | equal => simp [Fin.natAdd_zero, castLE_of_eq, mapTermRel] | rel => simp [Fin.natAdd_zero, castLE_of_eq, mapTermRel]; rfl - | imp _ _ ih1 ih2 => simp [mapTermRel, ih1, ih2] - | all _ ih3 => simp [mapTermRel, ih3, castLE] + | imp _ _ ih1 ih2 => simp_all [mapTermRel] + | all _ ih3 => simp_all [mapTermRel] /-- Substitutes the variables in a given formula with terms. -/ def subst {n : ℕ} (φ : L.BoundedFormula α n) (f : α → L.Term β) : L.BoundedFormula β n := diff --git a/Mathlib/NumberTheory/ADEInequality.lean b/Mathlib/NumberTheory/ADEInequality.lean index 4c41bee0d6a33..aabf630647bb3 100644 --- a/Mathlib/NumberTheory/ADEInequality.lean +++ b/Mathlib/NumberTheory/ADEInequality.lean @@ -148,7 +148,7 @@ theorem Admissible.one_lt_sumInv {pqr : Multiset ℕ+} : Admissible pqr → 1 < all_goals rw [← H, E', sumInv_pqr] conv_rhs => simp only [OfNat.ofNat, PNat.mk_coe] - rfl + norm_num theorem lt_three {p q r : ℕ+} (hpq : p ≤ q) (hqr : q ≤ r) (H : 1 < sumInv {p, q, r}) : p < 3 := by have h3 : (0 : ℚ) < 3 := by norm_num diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 772c5d12f2dff..7f3a1d190e6f7 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -121,6 +121,10 @@ theorem bertrand_main_inequality {n : ℕ} (n_large : 512 ≤ n) : · norm_num1 · exact cast_div_le.trans (by norm_cast) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- A lemma that tells us that, in the case where Bertrand's postulate does not hold, the prime factorization of the central binomial coefficient only has factors at most `2 * n / 3 + 1`. -/ @@ -134,7 +138,7 @@ theorem centralBinom_factorization_small (n : ℕ) (n_large : 2 < n) rw [Finset.mem_range, Nat.lt_succ_iff] at hx h2x rw [not_le, div_lt_iff_lt_mul' three_pos, mul_comm x] at h2x replace no_prime := not_exists.mp no_prime x - rw [← and_assoc, not_and', not_and_or, not_lt] at no_prime + rw [← _root_.and_assoc, not_and', not_and_or, not_lt] at no_prime cases' no_prime hx with h h · rw [factorization_eq_zero_of_non_prime n.centralBinom h, Nat.pow_zero] · rw [factorization_centralBinom_of_two_mul_self_lt_three_mul n_large h h2x, Nat.pow_zero] diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index e526d1422118e..c4b31c98a929a 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -168,7 +168,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( convert_to (discr K fun i : Fin 1 ↦ (algebraMap K L) (-1) ^ ↑i) = _ · congr ext i - simp only [map_neg, map_one, Function.comp_apply, Fin.coe_fin_one, _root_.pow_zero] + simp only [map_neg, map_one, Function.comp_apply, Fin.val_eq_zero, _root_.pow_zero] suffices (e.symm i : ℕ) = 0 by simp [this] rw [← Nat.lt_one_iff] convert (e.symm i).2 diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 0c98cefad6ecd..43de39c1d54b0 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -66,10 +66,14 @@ theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : theorem properDivisors.not_self_mem : ¬n ∈ properDivisors n := by simp [properDivisors] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_properDivisors {m : ℕ} : n ∈ properDivisors m ↔ n ∣ m ∧ n < m := by rcases eq_or_ne m 0 with (rfl | hm); · simp [properDivisors] - simp only [and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] + simp only [_root_.and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] theorem insert_self_properDivisors (h : n ≠ 0) : insert n (properDivisors n) = divisors n := by rw [divisors, properDivisors, Ico_succ_right_eq_insert_Ico (one_le_iff_ne_zero.2 h), @@ -96,11 +100,15 @@ theorem dvd_of_mem_divisors {m : ℕ} (h : n ∈ divisors m) : n ∣ m := by · apply dvd_zero · simp [mem_divisors.1 h] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} : x ∈ divisorsAntidiagonal n ↔ x.fst * x.snd = n ∧ n ≠ 0 := by simp only [divisorsAntidiagonal, Finset.mem_Ico, Ne, Finset.mem_filter, Finset.mem_product] - rw [and_comm] + rw [_root_.and_comm] apply and_congr_right rintro rfl constructor <;> intro h @@ -421,10 +429,14 @@ theorem sum_properDivisors_eq_one_iff_prime : ∑ x ∈ n.properDivisors, x = 1 (one_mem_properDivisors_iff_one_lt.2 (succ_lt_succ (Nat.succ_pos _)))) ((sum_singleton _ _).trans h.symm) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem mem_properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ} : x ∈ properDivisors (p ^ k) ↔ ∃ (j : ℕ) (_ : j < k), x = p ^ j := by rw [mem_properDivisors, Nat.dvd_prime_pow pp, ← exists_and_right] - simp only [exists_prop, and_assoc] + simp only [exists_prop, _root_.and_assoc] apply exists_congr intro a constructor <;> intro h @@ -469,13 +481,17 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ rw [← map_swap_divisorsAntidiagonal, Finset.prod_map] exact prod_divisorsAntidiagonal fun i j => f j i +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : n.primeFactors = (divisors n).filter Prime := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q - simpa [hn, hn.ne', mem_primeFactorsList] using and_comm + simpa [hn, hn.ne', mem_primeFactorsList] using _root_.and_comm @[deprecated (since := "2024-07-17")] alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index 392736e26084d..b2cb30daa0d3e 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -180,6 +180,10 @@ private theorem sum_Ico_eq_card_lt {p q : ℕ} : (by simp (config := { contextual := true }) only [mem_filter, mem_sigma, and_self_iff, forall_true_iff, mem_product]) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Each of the sums in this lemma is the cardinality of the set of integer points in each of the two triangles formed by the diagonal of the rectangle `(0, p/2) × (0, q/2)`. Adding them gives the number of points in the rectangle. -/ @@ -193,7 +197,7 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 card_equiv (Equiv.prodComm _ _) (fun ⟨_, _⟩ => by simp (config := { contextual := true }) only [mem_filter, and_self_iff, Prod.swap_prod_mk, - forall_true_iff, mem_product, Equiv.prodComm_apply, and_assoc, and_left_comm]) + forall_true_iff, mem_product, Equiv.prodComm_apply, _root_.and_assoc, and_left_comm]) have hdisj : Disjoint ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.2 * p ≤ x.1 * q) ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.1 * q ≤ x.2 * p) := by diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index 74dce799b0ee1..c4b313daf7c59 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -156,6 +156,10 @@ 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 @@ -165,7 +169,7 @@ theorem eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [NeZero b] : J(a | b) = 0 -- 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, - and_assoc, and_comm]) + _root_.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 @@ -214,7 +218,7 @@ theorem sq_one' {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a ^ 2 | b) = 1 := by r /-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ theorem mod_left (a : ℤ) (b : ℕ) : J(a | b) = J(a % b | b) := congr_arg List.prod <| - List.pmap_congr _ + List.pmap_congr_left _ (by -- Porting note: Lean does not synthesize the instance [Fact (Nat.Prime p)] automatically -- (it is needed for `legendreSym.mod` on line 227). Thus, we name the hypothesis @@ -309,7 +313,7 @@ theorem value_at (a : ℤ) {R : Type*} [CommSemiring R] (χ : R →* ℤ) conv_rhs => rw [← prod_primeFactorsList hb.pos.ne', cast_list_prod, map_list_prod χ] rw [jacobiSym, List.map_map, ← List.pmap_eq_map Nat.Prime _ _ fun _ => prime_of_mem_primeFactorsList] - congr 1; apply List.pmap_congr + congr 1; apply List.pmap_congr_left exact fun p h pp _ => hp p pp (hb.ne_two_of_dvd_nat <| dvd_of_mem_primeFactorsList h) /-- If `b` is odd, then `J(-1 | b)` is given by `χ₄ b`. -/ diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 699d1221574f7..12b15e1ddd7d7 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -217,7 +217,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : fun w => ⟨(w x) / 2, div_nonneg (AbsoluteValue.nonneg _ _) (by norm_num)⟩ suffices ∀ w, w ≠ w₁ → f w ≠ 0 by obtain ⟨g, h_geqf, h_gprod⟩ := adjust_f K B this - obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt (f := g) + obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt K (f := g) (by rw [convexBodyLT_volume]; convert hB; exact congr_arg ((↑) : NNReal → ENNReal) h_gprod) refine ⟨y, h_ynz, fun w hw => (h_geqf w hw ▸ h_yle w).trans ?_, ?_⟩ · rw [← Rat.cast_le (K := ℝ), Rat.cast_natCast] diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 8522d8eb5a00d..57ca7ff109082 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -179,7 +179,7 @@ theorem normalize : PythagoreanTriple (x / Int.gcd x y) (y / Int.gcd x y) (z / I have hz : z = 0 := by simpa only [PythagoreanTriple, hx, hy, add_zero, zero_eq_mul, mul_zero, or_self_iff] using h - simp only [hx, hy, hz, Int.zero_div] + simp only [hx, hy, hz] exact zero rcases h.gcd_dvd with ⟨z0, rfl⟩ obtain ⟨k, x0, y0, k0, h2, rfl, rfl⟩ : diff --git a/Mathlib/NumberTheory/SmoothNumbers.lean b/Mathlib/NumberTheory/SmoothNumbers.lean index 18f25f9900202..a130280fc2731 100644 --- a/Mathlib/NumberTheory/SmoothNumbers.lean +++ b/Mathlib/NumberTheory/SmoothNumbers.lean @@ -233,7 +233,7 @@ def equivProdNatFactoredNumbers {s : Finset ℕ} {p : ℕ} (hp : p.Prime) (hs : refine prod_eq <| (filter _ <| perm_primeFactorsList_mul (pow_ne_zero e hp.ne_zero) hm₀).trans ?_ rw [filter_append, hp.primeFactorsList_pow, - filter_eq_nil.mpr fun q hq ↦ by rw [mem_replicate] at hq; simp [hq.2, hs], + filter_eq_nil_iff.mpr fun q hq ↦ by rw [mem_replicate] at hq; simp [hq.2, hs], nil_append, filter_eq_self.mpr fun q hq ↦ by simp only [hm q hq, decide_True]] right_inv := by rintro ⟨m, hm₀, hm⟩ diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index 5d7a765adee86..a08b8c479c8ee 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -121,6 +121,10 @@ private theorem sum_four_squares_of_two_mul_sum_four_squares {m a b c d : ℤ} have : (∑ x, f (σ x) ^ 2) = ∑ x, f x ^ 2 := Equiv.sum_comp σ (f · ^ 2) simpa only [← hx, ← hy, Fin.sum_univ_four, add_assoc] using this +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Lagrange's **four squares theorem** for a prime number. Use `Nat.sum_four_squares` instead. -/ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : ∃ a b c d : ℕ, a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = p := by @@ -170,9 +174,10 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : -- The quotient `r` is not zero, because otherwise `f a = f b = f c = f d = 0`, hence -- `m` divides each `a`, `b`, `c`, `d`, thus `m ∣ p` which is impossible. rcases (zero_le r).eq_or_gt with rfl | hr₀ - · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [and_assoc] using hr + · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [_root_.and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d := by - simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] + simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, + _root_.and_self] have : m * m ∣ m * p := habcd ▸ ⟨a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2, by ring⟩ rw [mul_dvd_mul_iff_left hm₀.ne'] at this exact (hp.eq_one_or_self_of_dvd _ this).elim hm₁.ne' hmp.ne diff --git a/Mathlib/Order/Directed.lean b/Mathlib/Order/Directed.lean index 39d27f9b220b1..29285fca18124 100644 --- a/Mathlib/Order/Directed.lean +++ b/Mathlib/Order/Directed.lean @@ -167,7 +167,7 @@ instance OrderDual.isDirected_le [LE α] [IsDirected α (· ≥ ·)] : IsDirecte /-- A monotone function on an upwards-directed type is directed. -/ theorem directed_of_isDirected_le [LE α] [IsDirected α (· ≤ ·)] {f : α → β} {r : β → β → Prop} (H : ∀ ⦃i j⦄, i ≤ j → r (f i) (f j)) : Directed r f := - directed_id.mono_comp H + directed_id.mono_comp _ H theorem Monotone.directed_le [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {f : α → β} : Monotone f → Directed (· ≤ ·) f := diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index dfe2683487ad8..66d78953d6538 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -668,7 +668,7 @@ theorem HasBasis.eq_iInf (h : l.HasBasis (fun _ => True) s) : l = ⨅ i, 𝓟 (s theorem hasBasis_iInf_principal {s : ι → Set α} (h : Directed (· ≥ ·) s) [Nonempty ι] : (⨅ i, 𝓟 (s i)).HasBasis (fun _ => True) s := ⟨fun t => by - simpa only [true_and] using mem_iInf_of_directed (h.mono_comp monotone_principal.dual) t⟩ + simpa only [true_and] using mem_iInf_of_directed (h.mono_comp _ monotone_principal.dual) t⟩ /-- If `s : ι → Set α` is an indexed family of sets, then finite intersections of `s i` form a basis of `⨅ i, 𝓟 (s i)`. -/ @@ -683,7 +683,7 @@ theorem hasBasis_biInf_principal {s : β → Set α} {S : Set β} (h : DirectedO ⟨fun t => by refine mem_biInf_of_directed ?_ ne rw [directedOn_iff_directed, ← directed_comp] at h ⊢ - refine h.mono_comp ?_ + refine h.mono_comp _ ?_ exact fun _ _ => principal_mono.2⟩ theorem hasBasis_biInf_principal' {ι : Type*} {p : ι → Prop} {s : ι → Set α} diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 5f4bd78be7ae1..f77c690521b89 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -111,7 +111,8 @@ abbrev OrderHomClass (F : Type*) (α β : outParam Type*) [LE α] [LE β] [FunLi /-- `OrderIsoClass F α β` states that `F` is a type of order isomorphisms. You should extend this class when you extend `OrderIso`. -/ -class OrderIsoClass (F α β : Type*) [LE α] [LE β] [EquivLike F α β] : Prop where +class OrderIsoClass (F : Type*) (α β : outParam Type*) [LE α] [LE β] [EquivLike F α β] : + Prop where /-- An order isomorphism respects `≤`. -/ map_le_map_iff (f : F) {a b : α} : f a ≤ f b ↔ a ≤ b diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 971df11a1b4ff..92fa0698730b3 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -59,14 +59,16 @@ section /-- `TopHomClass F α β` states that `F` is a type of `⊤`-preserving morphisms. You should extend this class when you extend `TopHom`. -/ -class TopHomClass (F α β : Type*) [Top α] [Top β] [FunLike F α β] : Prop where +class TopHomClass (F : Type*) (α β : outParam Type*) [Top α] [Top β] [FunLike F α β] : + Prop where /-- A `TopHomClass` morphism preserves the top element. -/ map_top (f : F) : f ⊤ = ⊤ /-- `BotHomClass F α β` states that `F` is a type of `⊥`-preserving morphisms. You should extend this class when you extend `BotHom`. -/ -class BotHomClass (F α β : Type*) [Bot α] [Bot β] [FunLike F α β] : Prop where +class BotHomClass (F : Type*) (α β : outParam Type*) [Bot α] [Bot β] [FunLike F α β] : + Prop where /-- A `BotHomClass` morphism preserves the bottom element. -/ map_bot (f : F) : f ⊥ = ⊥ diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 133d9233ac466..d59c44eaa2139 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -185,10 +185,14 @@ theorem Ico_image_const_sub_eq_Ico (hac : a ≤ c) : rintro ⟨x, hx, rfl⟩ omega +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem Ico_succ_left_eq_erase_Ico : Ico a.succ b = erase (Ico a b) a := by ext x - rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← and_assoc, ne_comm, @and_comm (a ≠ x), - lt_iff_le_and_ne] + rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← _root_.and_assoc, ne_comm, + _root_.and_comm (a := a ≠ x), lt_iff_le_and_ne] theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := by induction' n with n ih diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 982bf1b0e6176..e801c4d599937 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -58,7 +58,7 @@ satisfy `r a b → s (f a) (f b)`. The relations `r` and `s` are `outParam`s since figuring them out from a goal is a higher-order matching problem that Lean usually can't do unaided. -/ -class RelHomClass (F : Type*) {α β : Type*} (r : outParam <| α → α → Prop) +class RelHomClass (F : Type*) {α β : outParam Type*} (r : outParam <| α → α → Prop) (s : outParam <| β → β → Prop) [FunLike F α β] : Prop where /-- A `RelHomClass` sends related elements to related elements -/ map_rel : ∀ (f : F) {a b}, r a b → s (f a) (f b) diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index b2e2b120664d7..f721b4e776979 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -185,7 +185,7 @@ end LinearLocallyFiniteOrder section toZ -- Requiring either of `IsSuccArchimedean` or `IsPredArchimedean` is equivalent. -variable [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} +variable [SuccOrder ι] [IsSuccArchimedean ι] [P : PredOrder ι] {i0 i : ι} -- For "to_Z" @@ -194,12 +194,13 @@ variable [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} the range of `toZ`. -/ def toZ (i0 i : ι) : ℤ := dite (i0 ≤ i) (fun hi ↦ Nat.find (exists_succ_iterate_of_le hi)) fun hi ↦ - -Nat.find (exists_pred_iterate_of_le (not_le.mp hi).le) + -Nat.find (exists_pred_iterate_of_le (α := ι) (not_le.mp hi).le) theorem toZ_of_ge (hi : i0 ≤ i) : toZ i0 i = Nat.find (exists_succ_iterate_of_le hi) := dif_pos hi -theorem toZ_of_lt (hi : i < i0) : toZ i0 i = -Nat.find (exists_pred_iterate_of_le hi.le) := +theorem toZ_of_lt (hi : i < i0) : + toZ i0 i = -Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ hi.le) := dif_neg (not_le.mpr hi) @[simp] @@ -310,7 +311,7 @@ theorem toZ_mono {i j : ι} (h_le : i ≤ j) : toZ i0 i ≤ toZ i0 j := by · exact le_of_not_le h · exact absurd h_le (not_le.mpr (hj.trans_le hi)) · exact (toZ_neg hi).le.trans (toZ_nonneg hj) - · let m := Nat.find (exists_pred_iterate_of_le h_le) + · let m := Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ h_le) have hm : pred^[m] j = i := Nat.find_spec (exists_pred_iterate_of_le h_le) have hj_eq : i = pred^[(-toZ i0 j).toNat + m] i0 := by rw [← hm, add_comm] diff --git a/Mathlib/Order/SymmDiff.lean b/Mathlib/Order/SymmDiff.lean index 8805a9c3184ad..fa8d23d3aa081 100644 --- a/Mathlib/Order/SymmDiff.lean +++ b/Mathlib/Order/SymmDiff.lean @@ -82,7 +82,7 @@ theorem symmDiff_eq_Xor' (p q : Prop) : p ∆ q = Xor' p q := @[simp] theorem bihimp_iff_iff {p q : Prop} : p ⇔ q ↔ (p ↔ q) := - (iff_iff_implies_and_implies _ _).symm.trans Iff.comm + iff_iff_implies_and_implies.symm.trans Iff.comm @[simp] theorem Bool.symmDiff_eq_xor : ∀ p q : Bool, p ∆ q = xor p q := by decide diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index ec56f812e2f47..877ecc41e7f9d 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -176,7 +176,8 @@ theorem toOuterMeasure_apply_eq_one_iff : p.toOuterMeasure s = 1 ↔ p.support (fun x => Set.indicator_apply_le fun _ => le_rfl) hsa · suffices ∀ (x) (_ : x ∉ s), p x = 0 from _root_.trans (tsum_congr - fun a => (Set.indicator_apply s p a).trans (ite_eq_left_iff.2 <| symm ∘ this a)) p.tsum_coe + fun a => (Set.indicator_apply s p a).trans + (ite_eq_left_iff.2 <| symm ∘ this a)) p.tsum_coe exact fun a ha => (p.apply_eq_zero_iff a).2 <| Set.not_mem_subset h ha @[simp] diff --git a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean index d137f3c66d02f..31a9cbfa903d1 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean @@ -67,9 +67,11 @@ theorem toOuterMeasure_pure_apply : (pure a).toOuterMeasure s = if a ∈ s then refine (toOuterMeasure_apply (pure a) s).trans ?_ split_ifs with ha · refine (tsum_congr fun b => ?_).trans (tsum_ite_eq a 1) - exact ite_eq_left_iff.2 fun hb => symm (ite_eq_right_iff.2 fun h => (hb <| h.symm ▸ ha).elim) + exact ite_eq_left_iff.2 fun hb => + symm (ite_eq_right_iff.2 fun h => (hb <| h.symm ▸ ha).elim) · refine (tsum_congr fun b => ?_).trans tsum_zero - exact ite_eq_right_iff.2 fun hb => ite_eq_right_iff.2 fun h => (ha <| h ▸ hb).elim + exact ite_eq_right_iff.2 fun hb => + ite_eq_right_iff.2 fun h => (ha <| h ▸ hb).elim variable [MeasurableSpace α] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index fbe67ec9e3438..2b83e2348e8b1 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -146,7 +146,7 @@ theorem dZero_comp_eq : dZero A ∘ₗ (zeroCochainsLequiv A) = oneCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 0 1 := by ext x y show A.ρ y (x default) - x default = _ + ({0} : Finset _).sum _ - simp_rw [Fin.coe_fin_one, zero_add, pow_one, neg_smul, one_smul, + simp_rw [Fin.val_eq_zero, zero_add, pow_one, neg_smul, one_smul, Finset.sum_singleton, sub_eq_add_neg] rcongr i <;> exact Fin.elim0 i diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index 7532a1f66cbdf..bba7f1e3e7457 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -353,11 +353,11 @@ noncomputable instance instField [Fact (Irreducible f)] : Field (AdjoinRoot f) w ratCast_def q := by rw [← map_natCast (of f), ← map_intCast (of f), ← map_div₀, ← Rat.cast_def]; rfl nnqsmul_def q x := - AdjoinRoot.induction_on (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by + AdjoinRoot.induction_on f (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by simp only [smul_mk, of, RingHom.comp_apply, ← (mk f).map_mul, Polynomial.nnqsmul_eq_C_mul] qsmul_def q x := -- Porting note: I gave the explicit motive and changed `rw` to `simp`. - AdjoinRoot.induction_on (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by + AdjoinRoot.induction_on f (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by simp only [smul_mk, of, RingHom.comp_apply, ← (mk f).map_mul, Polynomial.qsmul_eq_C_mul] theorem coe_injective (h : degree f ≠ 0) : Function.Injective ((↑) : K → AdjoinRoot f) := diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 0a9301168ce1c..146fc3d8f12a9 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -395,14 +395,14 @@ variable [Algebra R A] [Algebra R A'] [IsScalarTower R A B] [IsScalarTower R A' /-- Integral closures are all isomorphic to each other. -/ noncomputable def equiv : A ≃ₐ[R] A' := AlgEquiv.ofAlgHom - (lift _ B (isIntegral := isIntegral_algebra R B)) - (lift _ B (isIntegral := isIntegral_algebra R B)) + (lift R A' B (isIntegral := isIntegral_algebra R B)) + (lift R A B (isIntegral := isIntegral_algebra R B)) (by ext x; apply algebraMap_injective A' R B; simp) (by ext x; apply algebraMap_injective A R B; simp) @[simp] theorem algebraMap_equiv (x : A) : algebraMap A' B (equiv R A B A' x) = algebraMap A B x := - algebraMap_lift A' B (isIntegral := isIntegral_algebra R B) x + algebraMap_lift R A' B (isIntegral := isIntegral_algebra R B) x end Equiv diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 838c66b8f3863..332bccee7c8f5 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -178,19 +178,21 @@ theorem free_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] Module.finrank_finsupp_self, one_mul, add_zero] rw [Module.finrank_eq_card_chooseBasisIndex] -- On the other hand, `m ⊗ M → M` injective => `Tor₁(k, M) = 0` => `k ⊗ ker(i) → kᴵ` injective. - have := @lTensor_injective_of_exact_of_exact_of_rTensor_injective - (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) - (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) inferInstance inferInstance inferInstance - inferInstance inferInstance inferInstance intro x - apply @this (LinearMap.ker i).subtype i (LinearMap.exact_subtype_mkQ 𝔪) - (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H - (Module.Flat.lTensor_preserves_injective_linearMap _ Subtype.val_injective) - apply hi'.injective - rw [LinearMap.baseChange_eq_ltensor] - erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] - rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] - simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] + refine lTensor_injective_of_exact_of_exact_of_rTensor_injective + (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) + (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) + (g₁ := (LinearMap.ker i).subtype) (g₂ := i) (LinearMap.exact_subtype_mkQ 𝔪) + (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H ?_ ?_ + · apply Module.Flat.lTensor_preserves_injective_linearMap + (N := LinearMap.ker i) (N' := I →₀ R) + (L := (LinearMap.ker i).subtype) + exact Subtype.val_injective + · apply hi'.injective + rw [LinearMap.baseChange_eq_ltensor] + erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] + rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] + simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] -- TODO: Generalise this to finite free modules. theorem free_of_flat_of_localRing [Module.FinitePresentation R P] [Module.Flat R P] : diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 1960bbcd7e01d..f4b231a725789 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -132,7 +132,7 @@ noncomputable abbrev toField : Field K where lemma surjective_iff_isField [IsDomain R] : Function.Surjective (algebraMap R K) ↔ IsField R where mp h := (RingEquiv.ofBijective (algebraMap R K) - ⟨IsFractionRing.injective R K, h⟩).toMulEquiv.isField (IsFractionRing.toField R).toIsField + ⟨IsFractionRing.injective R K, h⟩).toMulEquiv.isField _ (IsFractionRing.toField R).toIsField mpr h := letI := h.toField (IsLocalization.atUnits R _ (S := K) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index 3aa09ead7b4ab..ddadb79adcf83 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -250,13 +250,17 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] + simp only [Set.Ioo, Set.mem_setOf_eq, _root_.and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 55bb7d2fd0d44..891e01b70140c 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -29,7 +29,7 @@ variable {R : Type u} {S : Type v} {T : Type w} [NonUnitalNonAssocSemiring R] (M /-- `NonUnitalSubsemiringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both an additive submonoid and also a multiplicative subsemigroup. -/ -class NonUnitalSubsemiringClass (S : Type*) (R : Type u) [NonUnitalNonAssocSemiring R] +class NonUnitalSubsemiringClass (S : Type*) (R : outParam (Type u)) [NonUnitalNonAssocSemiring R] [SetLike S R] extends AddSubmonoidClass S R : Prop where mul_mem : ∀ {s : S} {a b : R}, a ∈ s → b ∈ s → a * b ∈ s diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index 4a1ad15664ee8..43e050ec7489d 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -104,12 +104,17 @@ theorem cyclotomic_pos {n : ℕ} (hn : 2 < n) {R} [LinearOrderedCommRing R] (x : exact hn'.ne' hi.2.2.1 · simpa only [eval_X, eval_one, cyclotomic_two, eval_add] using h.right.le +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem cyclotomic_pos_and_nonneg (n : ℕ) {R} [LinearOrderedCommRing R] (x : R) : (1 < x → 0 < eval x (cyclotomic n R)) ∧ (1 ≤ x → 0 ≤ eval x (cyclotomic n R)) := by rcases n with (_ | _ | _ | n) - · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, and_self] + · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, + _root_.and_self] · simp only [zero_add, cyclotomic_one, eval_sub, eval_X, eval_one, sub_pos, imp_self, sub_nonneg, - and_self] + _root_.and_self] · simp only [zero_add, reduceAdd, cyclotomic_two, eval_add, eval_X, eval_one] constructor <;> intro <;> linarith · constructor <;> intro <;> [skip; apply le_of_lt] <;> apply cyclotomic_pos (by omega) diff --git a/Mathlib/RingTheory/RingHomProperties.lean b/Mathlib/RingTheory/RingHomProperties.lean index 1b12d138c4092..7008cd0ef9a74 100644 --- a/Mathlib/RingTheory/RingHomProperties.lean +++ b/Mathlib/RingTheory/RingHomProperties.lean @@ -185,9 +185,11 @@ lemma toMorphismProperty_respectsIso_iff : · intro X Y Z e f hf exact h.left f e.commRingCatIsoToRingEquiv hf · intro X Y Z _ _ _ f e hf - exact h.postcomp e.toCommRingCatIso.hom (CommRingCat.ofHom f) hf + exact MorphismProperty.RespectsIso.postcomp (toMorphismProperty P) + e.toCommRingCatIso.hom (CommRingCat.ofHom f) hf · intro X Y Z _ _ _ f e - exact h.precomp e.toCommRingCatIso.hom (CommRingCat.ofHom f) + exact MorphismProperty.RespectsIso.precomp (toMorphismProperty P) + e.toCommRingCatIso.hom (CommRingCat.ofHom f) end ToMorphismProperty diff --git a/Mathlib/RingTheory/SimpleRing/Basic.lean b/Mathlib/RingTheory/SimpleRing/Basic.lean index c9f5a875cc0b8..3f9711cafff02 100644 --- a/Mathlib/RingTheory/SimpleRing/Basic.lean +++ b/Mathlib/RingTheory/SimpleRing/Basic.lean @@ -82,6 +82,6 @@ lemma isField_center (A : Type*) [Ring A] [IsSimpleRing A] : IsField (Subring.ce end IsSimpleRing lemma isSimpleRing_iff_isField (A : Type*) [CommRing A] : IsSimpleRing A ↔ IsField A := - ⟨fun _ ↦ Subring.topEquiv.symm.toMulEquiv.isField <| by + ⟨fun _ ↦ Subring.topEquiv.symm.toMulEquiv.isField _ <| by rw [← Subring.center_eq_top A]; exact IsSimpleRing.isField_center A, fun h ↦ letI := h.toField; inferInstance⟩ diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index f9ed2c1090626..988a61b9c879e 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -544,14 +544,14 @@ instance IsStandardSmooth.baseChange [IsStandardSmooth.{t, w} R S] : IsStandardSmooth.{t, w} T (T ⊗[R] S) where out := by obtain ⟨⟨P⟩⟩ := ‹IsStandardSmooth R S› - exact ⟨P.baseChange T⟩ + exact ⟨P.baseChange R S T⟩ instance IsStandardSmoothOfRelativeDimension.baseChange [IsStandardSmoothOfRelativeDimension.{t, w} n R S] : IsStandardSmoothOfRelativeDimension.{t, w} n T (T ⊗[R] S) where out := by obtain ⟨P, hP⟩ := ‹IsStandardSmoothOfRelativeDimension n R S› - exact ⟨P.baseChange T, hP⟩ + exact ⟨P.baseChange R S T, hP⟩ end BaseChange diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 1cd7eb9c7c18a..08b850cedaa5a 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -95,15 +95,19 @@ 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 - rw [and_comm, Bijective, and_congr_left_iff] + rw [_root_.and_comm, Bijective, and_congr_left_iff] intro h have := Fintype.ofFinite α have := Fintype.ofSurjective f h revert h - rw [← and_congr_left_iff, ← Bijective, ← and_comm, + rw [← and_congr_left_iff, ← Bijective, ← _root_.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 : α → β} diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 39157d0e58cab..51dcb46a44da0 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -842,7 +842,7 @@ def dbgTraceACState : CCM Unit := do def mkACProof (e₁ e₂ : Expr) : MetaM Expr := do let eq ← mkEq e₁ e₂ let .mvar m ← mkFreshExprSyntheticOpaqueMVar eq | failure - AC.rewriteUnnormalized m + AC.rewriteUnnormalizedRefl m let pr ← instantiateMVars (.mvar m) mkExpectedTypeHint pr eq @@ -1470,7 +1470,8 @@ partial def propagateEqUp (e : Expr) : CCM Unit := do if ← isInterpretedValue ra <&&> isInterpretedValue rb <&&> pure (ra.int?.isNone || ra.int? != rb.int?) then raNeRb := some - (Expr.app (.proj ``Iff 0 (← mkAppM ``bne_iff_ne #[ra, rb])) (← mkEqRefl (.const ``true []))) + (Expr.app (.proj ``Iff 0 (← mkAppOptM ``bne_iff_ne #[none, none, none, ra, rb])) + (← mkEqRefl (.const ``true []))) else if let some c₁ ← isConstructorApp? ra then if let some c₂ ← isConstructorApp? rb then @@ -1808,7 +1809,8 @@ def propagateValueInconsistency (e₁ e₂ : Expr) : CCM Unit := do let some eqProof ← getEqProof e₁ e₂ | failure let trueEqFalse ← mkEq (.const ``True []) (.const ``False []) let neProof := - Expr.app (.proj ``Iff 0 (← mkAppM ``bne_iff_ne #[e₁, e₂])) (← mkEqRefl (.const ``true [])) + Expr.app (.proj ``Iff 0 (← mkAppOptM ``bne_iff_ne #[none, none, none, e₁, e₂])) + (← mkEqRefl (.const ``true [])) let H ← mkAbsurd trueEqFalse eqProof neProof pushEq (.const ``True []) (.const ``False []) H diff --git a/Mathlib/Tactic/IntervalCases.lean b/Mathlib/Tactic/IntervalCases.lean index 52169b6406e4c..37e90dfc742f4 100644 --- a/Mathlib/Tactic/IntervalCases.lean +++ b/Mathlib/Tactic/IntervalCases.lean @@ -5,6 +5,7 @@ Authors: Kim Morrison, Mario Carneiro -/ import Mathlib.Tactic.NormNum import Mathlib.Tactic.FinCases +import Mathlib.Control.Basic /-! # Case bash on variables in finite intervals diff --git a/Mathlib/Tactic/NormNum/DivMod.lean b/Mathlib/Tactic/NormNum/DivMod.lean index 3037be0ca7da7..821532050c549 100644 --- a/Mathlib/Tactic/NormNum/DivMod.lean +++ b/Mathlib/Tactic/NormNum/DivMod.lean @@ -147,8 +147,8 @@ theorem isInt_dvd_true : {a b : ℤ} → {a' b' c : ℤ} → | _, _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, rfl => ⟨_, rfl⟩ theorem isInt_dvd_false : {a b : ℤ} → {a' b' : ℤ} → - IsInt a a' → IsInt b b' → Int.mod b' a' != 0 → ¬a ∣ b - | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, e => mt Int.mod_eq_zero_of_dvd (by simpa using e) + IsInt a a' → IsInt b b' → Int.emod b' a' != 0 → ¬a ∣ b + | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, e => mt Int.emod_eq_zero_of_dvd (by simpa using e) /-- The `norm_num` extension which identifies expressions of the form `(a : ℤ) ∣ b`, such that `norm_num` successfully recognises both `a` and `b`. -/ @@ -167,7 +167,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ haveI' : Int.mul $na $c =Q $nb := ⟨⟩ return .isTrue q(isInt_dvd_true $pa $pb (.refl $nb)) else - have : Q(Int.mod $nb $na != 0) := (q(Eq.refl true) : Expr) + have : Q(Int.emod $nb $na != 0) := (q(Eq.refl true) : Expr) return .isFalse q(isInt_dvd_false $pa $pb $this) end Mathlib.Meta.NormNum diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 498f20124180d..dcd4f4ddf5f6f 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -1014,7 +1014,7 @@ def ExSum.evalInv {a : Q($α)} (czα : Option Q(CharZero $α)) (va : ExSum sα a match va with | ExSum.zero => pure ⟨_, .zero, (q(inv_zero (R := $α)) : Expr)⟩ | ExSum.add va ExSum.zero => do - let ⟨_, vb, pb⟩ ← va.evalInv dα czα + let ⟨_, vb, pb⟩ ← va.evalInv sα dα czα pure ⟨_, vb.toSum, (q(inv_single $pb) : Expr)⟩ | va => do let ⟨_, vb, pb⟩ ← evalInvAtom sα dα a diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index c996e5552b6df..461949d04ab7b 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -86,7 +86,8 @@ def _root_.Lean.MVarId.wlog (goal : MVarId) (h : Option Name) (P : Expr) let hGoal := HExpr.mvarId! /- Begin the "reduction goal" which will contain hypotheses `H` and `¬h`. For now, it only contains `H`. Keep track of that hypothesis' FVarId. -/ - let (HFVarId, reductionGoal) ← goal.assertHypotheses #[⟨H, HType, HExpr⟩] + let (HFVarId, reductionGoal) ← + goal.assertHypotheses #[{ userName := H, type := HType, value := HExpr }] let HFVarId := HFVarId[0]! /- Clear the reverted fvars from the branch that will contain `h` as a hypothesis. -/ let hGoal ← hGoal.tryClearMany revertedFVars diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 12325734b9e55..55ef7dc3e20fa 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -363,7 +363,7 @@ theorem applyId_injective [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup · symm; rw [h] rw [← List.applyId_zip_eq] <;> assumption · rw [← h₁.length_eq] - rw [List.getElem?_eq_some] at hx + rw [List.getElem?_eq_some_iff] at hx cases' hx with hx hx' exact hx · rw [← applyId_mem_iff h₀ h₁] at hx hy diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index dbe91b1484df4..2c9ce7c08f2c8 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -102,7 +102,7 @@ def toNormedField : NormedField L := · set δ : ℝ≥0 := hv.hom ε with hδ have hδ_pos : 0 < δ := by rw [hδ, ← _root_.map_zero hv.hom] - exact hv.strictMono (Units.zero_lt ε) + exact hv.strictMono _ (Units.zero_lt ε) use δ, hδ_pos apply subset_trans _ hε intro x hx diff --git a/Mathlib/Topology/Category/LightProfinite/Extend.lean b/Mathlib/Topology/Category/LightProfinite/Extend.lean index bed0658153636..4f7e5eb315b90 100644 --- a/Mathlib/Topology/Category/LightProfinite/Extend.lean +++ b/Mathlib/Topology/Category/LightProfinite/Extend.lean @@ -118,7 +118,7 @@ then `cone G c.pt` is a limit cone. -/ noncomputable def isLimitCone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsLimit <| G.mapCone c) : - IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ hc' + IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ _ hc' end Limit @@ -158,7 +158,7 @@ noncomputable def isColimitCocone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsColimit <| G.mapCocone c.op) : IsColimit (cocone G c.pt) := haveI := functorOp_final c hc - (Functor.final_comp (opOpEquivalence ℕ).functor (functorOp c)).isColimitWhiskerEquiv _ hc' + (Functor.final_comp (opOpEquivalence ℕ).functor (functorOp c)).isColimitWhiskerEquiv _ _ hc' end Colimit diff --git a/Mathlib/Topology/Category/Profinite/Extend.lean b/Mathlib/Topology/Category/Profinite/Extend.lean index 8db0cc4dd84e9..2c00a5e5808e0 100644 --- a/Mathlib/Topology/Category/Profinite/Extend.lean +++ b/Mathlib/Topology/Category/Profinite/Extend.lean @@ -141,7 +141,7 @@ then `cone G c.pt` is a limit cone. -/ noncomputable def isLimitCone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsLimit <| G.mapCone c) : - IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ hc' + IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ _ hc' end Limit @@ -177,7 +177,7 @@ are epimorphic, then `cocone G c.pt` is a colimit cone. -/ noncomputable def isColimitCocone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsColimit <| G.mapCocone c.op) : - IsColimit (cocone G c.pt) := (functorOp_final c hc).isColimitWhiskerEquiv _ hc' + IsColimit (cocone G c.pt) := (functorOp_final c hc).isColimitWhiskerEquiv _ _ hc' end Colimit diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index e77f01dabfc8e..34583c8f81047 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -39,8 +39,8 @@ section /-- `ContinuousMapClass F α β` states that `F` is a type of continuous maps. You should extend this class when you extend `ContinuousMap`. -/ -class ContinuousMapClass (F α β : Type*) [TopologicalSpace α] [TopologicalSpace β] - [FunLike F α β] : Prop where +class ContinuousMapClass (F : Type*) (α β : outParam Type*) + [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] : Prop where /-- Continuity -/ map_continuous (f : F) : Continuous f diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index b94867562a459..0741095653d32 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -66,7 +66,7 @@ infixl:25 " →ᵈ " => Dilation /-- `DilationClass F α β r` states that `F` is a type of `r`-dilations. You should extend this typeclass when you extend `Dilation`. -/ -class DilationClass (F α β : Type*) [PseudoEMetricSpace α] [PseudoEMetricSpace β] +class DilationClass (F : Type*) (α β : outParam Type*) [PseudoEMetricSpace α] [PseudoEMetricSpace β] [FunLike F α β] : Prop where edist_eq' : ∀ f : F, ∃ r : ℝ≥0, r ≠ 0 ∧ ∀ x y : α, edist (f x) (f y) = r * edist x y diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index efdf16b745cec..3dc4176d75c21 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -158,8 +158,8 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x ← Option.coe_def, Option.toList_some, take_append_of_le_length hMl.le, getElem_cons_succ] · exact single_le_sum (fun x _ => zero_le x) _ (mem_iff_get.2 ⟨⟨M, hM_lt⟩, getElem_zipWith⟩) · rcases hMl.eq_or_lt with (rfl | hMl) - · simp only [getElem_append_right' le_rfl, sub_self, getElem_singleton, dist_self, zero_le] - rw [getElem_append _ hMl] + · simp only [getElem_append_right le_rfl, sub_self, getElem_singleton, dist_self, zero_le] + rw [getElem_append_left hMl] have hlen : length (drop (M + 1) l) = length l - (M + 1) := length_drop _ _ have hlen_lt : length l - (M + 1) < length l := Nat.sub_lt_of_pos_le M.succ_pos hMl refine (ihn _ hlen_lt _ y _ hlen).trans ?_ diff --git a/Mathlib/Topology/Order/LawsonTopology.lean b/Mathlib/Topology/Order/LawsonTopology.lean index 4c730963736e1..61c08e2a6fe21 100644 --- a/Mathlib/Topology/Order/LawsonTopology.lean +++ b/Mathlib/Topology/Order/LawsonTopology.lean @@ -144,7 +144,7 @@ instance instIsLawson : IsLawson (WithLawson α) := ⟨rfl⟩ /-- If `α` is equipped with the Lawson topology, then it is homeomorphic to `WithLawson α`. -/ def homeomorph [TopologicalSpace α] [IsLawson α] : WithLawson α ≃ₜ α := - ofLawson.toHomeomorphOfInducing ⟨by erw [@IsLawson.topology_eq_lawson α _ _, induced_id]; rfl⟩ + ofLawson.toHomeomorphOfInducing ⟨by erw [IsLawson.topology_eq_lawson (α := α), induced_id]; rfl⟩ theorem isOpen_preimage_ofLawson {S : Set α} : IsOpen (ofLawson ⁻¹' S) ↔ (lawson α).IsOpen S := Iff.rfl diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index a1a7590a0efa4..b2be7ed715fd7 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -563,7 +563,7 @@ theorem TotallyBounded.image [UniformSpace β] {f : α → β} {s : Set α} (hs simp only [mem_image, iUnion_exists, biUnion_and', iUnion_iUnion_eq_right, image_subset_iff, preimage_iUnion, preimage_setOf_eq] simp? [subset_def] at hct says - simp only [mem_setOf_eq, subset_def, mem_iUnion, exists_prop', nonempty_prop] at hct + simp only [mem_setOf_eq, subset_def, mem_iUnion, exists_prop] at hct intro x hx simpa using hct x hx⟩ diff --git a/Mathlib/Util/CountHeartbeats.lean b/Mathlib/Util/CountHeartbeats.lean index 3b30d85e1aa29..bbc00430ba743 100644 --- a/Mathlib/Util/CountHeartbeats.lean +++ b/Mathlib/Util/CountHeartbeats.lean @@ -43,8 +43,8 @@ def runTacForHeartbeats (tac : TSyntax `Lean.Parser.Tactic.tacticSeq) (revert : Given a `List Nat`, return the minimum, maximum, and standard deviation. -/ def variation (counts : List Nat) : List Nat := - let min := counts.minimum?.getD 0 - let max := counts.maximum?.getD 0 + let min := counts.min?.getD 0 + let max := counts.max?.getD 0 let toFloat (n : Nat) := n.toUInt64.toFloat let toNat (f : Float) := f.toUInt64.toNat let counts' := counts.map toFloat diff --git a/lake-manifest.json b/lake-manifest.json index 027d26b4ddaa2..3f76a781eeec7 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f274aed7ae8d1addd3e70adaf3183ccc6e1ed43d", + "rev": "63c1c38b123b0741b7b7fd56fb8510f95bfd0e55", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ff420521a0c098891f4f44ecda9dd7ff57b50bad", + "rev": "b20a88676fd00affb90cbc9f1ff004ae588103b3", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "e285a7ade149c551c17a4b24f127e1ef782e4bb1", + "rev": "6b6ad220389444229d6b29c386b039e18345a003", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index 87bbbfd1f21f0..212047a97ac35 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -35,7 +35,8 @@ abbrev mathlibOnlyLinters : Array LeanOption := #[ ⟨`linter.style.longLine, true⟩, ⟨`linter.style.longFile, .ofNat 1500⟩, ⟨`linter.style.missingEnd, true⟩, - ⟨`linter.style.setOption, true⟩ + ⟨`linter.style.setOption, true⟩, + ⟨`aesop.warn.applyIff, false⟩ -- This became a problem after https://github.com/leanprover-community/aesop/commit/29cf094e84ae9852f0011b47b6ddc684ffe4be5f ] /-- These options are passed as `leanOptions` to building mathlib, as well as the diff --git a/lean-toolchain b/lean-toolchain index 89985206aca4e..7c79e97102467 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0 +leanprover/lean4:v4.13.0-rc2 diff --git a/scripts/noshake.json b/scripts/noshake.json index 735d7aacd17cf..648722b370159 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -360,7 +360,8 @@ ["Batteries.Data.Nat.Lemmas", "Mathlib.Data.List.Basic"], "Mathlib.Data.List.Lemmas": ["Mathlib.Data.List.InsertNth"], "Mathlib.Data.List.Defs": ["Batteries.Data.RBMap.Basic"], - "Mathlib.Data.List.Basic": ["Mathlib.Data.Option.Basic"], + "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"], diff --git a/test/ValuedCSP.lean b/test/ValuedCSP.lean index f60e3aed3e421..48b89612eacdc 100644 --- a/test/ValuedCSP.lean +++ b/test/ValuedCSP.lean @@ -40,6 +40,8 @@ private def exampleFiniteValuedInstance : exampleFiniteValuedCSP.Instance (Fin 2 example : exampleFiniteValuedInstance.IsOptimumSolution ![(0 : ℚ), (0 : ℚ)] := by intro s convert_to 0 ≤ exampleFiniteValuedInstance.evalSolution s + · simp [exampleFiniteValuedInstance, ValuedCSP.Instance.evalSolution] + exact Rat.zero_add 0 rw [ValuedCSP.Instance.evalSolution, exampleFiniteValuedInstance] convert_to 0 ≤ |s 0| + |s 1| · simp [ValuedCSP.unaryTerm, ValuedCSP.Term.evalSolution, Function.OfArity.uncurry] diff --git a/test/aesop_cat.lean b/test/aesop_cat.lean index 7790b9c1b983c..549cd69b92adb 100644 --- a/test/aesop_cat.lean +++ b/test/aesop_cat.lean @@ -10,6 +10,8 @@ example : Foo where x := sorry /-- +error: could not synthesize default value for field 'w' of 'Foo' using tactics +--- error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. Initial goal: ⊢ 35 = 37 diff --git a/test/matrix.lean b/test/matrix.lean index 962e193b1ef97..522acbe923c7d 100644 --- a/test/matrix.lean +++ b/test/matrix.lean @@ -154,7 +154,7 @@ example {α : Type _} [CommRing α] {a b c d : α} : Fin.isValue, of_apply, cons_val', empty_val', cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, head_cons, - Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, cons_val_succ, neg_mul, + Finset.univ_unique, Fin.val_succ, Fin.val_eq_zero, zero_add, pow_one, cons_val_succ, neg_mul, Fin.succ_succAbove_zero, Finset.sum_const, Finset.card_singleton, smul_neg, one_smul] ring @@ -167,7 +167,7 @@ example {α : Type _} [CommRing α] {a b c d e f g h i : α} : submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_cons, submatrix_submatrix, det_unique, Fin.default_eq_zero, Function.comp_apply, Fin.succ_one_eq_two, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, - Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, + Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.val_eq_zero, zero_add, pow_one, neg_mul, Fin.succ_succAbove_zero, Finset.sum_neg_distrib, Finset.sum_singleton, cons_val_succ, Fin.succ_succAbove_one, even_two, Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, one_smul] diff --git a/test/says.lean b/test/says.lean index e2f9a3f834662..4315e51cb870d 100644 --- a/test/says.lean +++ b/test/says.lean @@ -101,7 +101,8 @@ def very_long_lemma_name_aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa : Q → P := fun _ @[simp] def very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb : Q := trivial /-- -info: Try this: aesop? says simp_all only [very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb, +info: Try this: aesop? says + simp_all only [very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb, very_long_lemma_name_aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa] -/ #guard_msgs in From 435c3d7b1429d7ec31a3f98373e2b7d7cec47d22 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Fri, 4 Oct 2024 03:58:07 +0000 Subject: [PATCH 166/174] chore(Finset): remove old-style spellings (#17327) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Deprecate `sdiff_singleton_eq_self` in favour of `erase_eq_of_not_mem`: the spelling `s.erase a` is preferred over `s \ {a}`. Deprecate `sdiff_eq_self` in favour of `sdiff_eq_self_iff_disjoint`: the spelling `Disjoint s t` is preferred over `s ∩ t ⊆ ∅`. Deprecate `filter_inter_filter_neg_eq` in favour of `disjoint_filter_filter_neg`: the spelling `Disjoint s t` is preferred over `s ∩ t = ∅`. Amend statement of `filter_cons` to be simpler and match `filter_insert` more closely. --- Mathlib/Data/FinEnum.lean | 40 ++++++++-------------------------- Mathlib/Data/Finset/Basic.lean | 32 ++++++++++----------------- Mathlib/Data/Fintype/Fin.lean | 13 +++++------ 3 files changed, 27 insertions(+), 58 deletions(-) diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index bfb3c8e6783ba..c2d94fcfa23ad 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -107,41 +107,19 @@ def Finset.enum [DecidableEq α] : List α → List (Finset α) | [] => [∅] | x :: xs => do let r ← Finset.enum xs - [r, {x} ∪ r] + [r, insert x r] @[simp] theorem Finset.mem_enum [DecidableEq α] (s : Finset α) (xs : List α) : s ∈ Finset.enum xs ↔ ∀ x ∈ s, x ∈ xs := by - induction' xs with xs_hd generalizing s <;> simp [*, Finset.enum] - · simp [Finset.eq_empty_iff_forall_not_mem] - · constructor - · rintro ⟨a, h, h'⟩ x hx - cases' h' with _ h' a b - · right - apply h - subst a - exact hx - · simp only [h', mem_union, mem_singleton] at hx ⊢ - cases' hx with hx hx' - · exact Or.inl hx - · exact Or.inr (h _ hx') - · intro h - exists s \ ({xs_hd} : Finset α) - simp only [and_imp, mem_sdiff, mem_singleton] - simp only [or_iff_not_imp_left] at h - exists h - by_cases h : xs_hd ∈ s - · have : {xs_hd} ⊆ s := by - simp only [HasSubset.Subset, *, forall_eq, mem_singleton] - simp only [union_sdiff_of_subset this, or_true, Finset.union_sdiff_of_subset, - eq_self_iff_true] - · left - symm - simp only [sdiff_eq_self] - intro a - simp only [and_imp, mem_inter, mem_singleton] - rintro h₀ rfl - exact (h h₀).elim + induction xs generalizing s with + | nil => simp [enum, eq_empty_iff_forall_not_mem] + | cons x xs ih => + simp only [enum, List.bind_eq_bind, List.mem_bind, List.mem_cons, List.mem_singleton, + List.not_mem_nil, or_false, ih] + refine ⟨by aesop, fun hs => ⟨s.erase x, ?_⟩⟩ + simp only [or_iff_not_imp_left] at hs + simp (config := { contextual := true }) [eq_comm (a := s), or_iff_not_imp_left, hs] instance Finset.finEnum [FinEnum α] : FinEnum (Finset α) := ofList (Finset.enum (toList α)) (by intro; simp) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index b84e46a5f7a05..758307ad5bb02 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1916,7 +1916,7 @@ theorem union_sdiff_self (s t : Finset α) : (s ∪ t) \ t = s \ t := -- TODO: Do we want to delete this lemma and `Finset.disjUnion_singleton`, -- or instead add `Finset.union_singleton`/`Finset.singleton_union`? -theorem sdiff_singleton_eq_erase (a : α) (s : Finset α) : s \ singleton a = erase s a := by +theorem sdiff_singleton_eq_erase (a : α) (s : Finset α) : s \ {a} = erase s a := by ext rw [mem_erase, mem_sdiff, mem_singleton, and_comm] @@ -1972,9 +1972,9 @@ theorem erase_union_of_mem (ha : a ∈ t) (s : Finset α) : s.erase a ∪ t = s theorem union_erase_of_mem (ha : a ∈ s) (t : Finset α) : s ∪ t.erase a = s ∪ t := by rw [← insert_erase (mem_union_left t ha), erase_union_distrib, ← insert_union, insert_erase ha] -@[simp] -theorem sdiff_singleton_eq_self (ha : a ∉ s) : s \ {a} = s := - sdiff_eq_self_iff_disjoint.2 <| by simp [ha] +@[simp, deprecated erase_eq_of_not_mem (since := "2024-10-01")] +theorem sdiff_singleton_eq_self (ha : a ∉ s) : s \ {a} = s := by + rw [← erase_eq, erase_eq_of_not_mem ha] theorem Nontrivial.sdiff_singleton_nonempty {c : α} {s : Finset α} (hS : s.Nontrivial) : (s \ {c}).Nonempty := by @@ -2038,6 +2038,8 @@ theorem disjoint_sdiff_inter (s t : Finset α) : Disjoint (s \ t) (s ∩ t) := theorem sdiff_eq_self_iff_disjoint : s \ t = s ↔ Disjoint s t := sdiff_eq_self_iff_disjoint' +@[deprecated (since := "2024-10-01")] alias sdiff_eq_self := sdiff_eq_self_iff_disjoint + theorem sdiff_eq_self_of_disjoint (h : Disjoint s t) : s \ t = s := sdiff_eq_self_iff_disjoint.2 h @@ -2315,7 +2317,7 @@ theorem filter_singleton (a : α) : filter p {a} = if p a then {a} else ∅ := b split_ifs with h <;> by_cases h' : x = a <;> simp [h, h'] theorem filter_cons_of_pos (a : α) (s : Finset α) (ha : a ∉ s) (hp : p a) : - filter p (cons a s ha) = cons a (filter p s) (mem_filter.not.mpr <| mt And.left ha) := + filter p (cons a s ha) = cons a (filter p s) ((mem_of_mem_filter _).mt ha) := eq_of_veq <| Multiset.filter_cons_of_pos s.val hp theorem filter_cons_of_neg (a : α) (s : Finset α) (ha : a ∉ s) (hp : ¬p a) : @@ -2344,6 +2346,8 @@ theorem disjoint_filter_filter_neg (s t : Finset α) (p : α → Prop) Disjoint (s.filter p) (t.filter fun a => ¬p a) := disjoint_filter_filter' s t disjoint_compl_right +@[deprecated (since := "2024-10-01")] alias filter_inter_filter_neg_eq := disjoint_filter_filter_neg + theorem filter_disj_union (s : Finset α) (t : Finset α) (h : Disjoint s t) : filter p (disjUnion s t h) = (filter p s).disjUnion (filter p t) (disjoint_filter_filter h) := eq_of_veq <| Multiset.filter_add _ _ _ @@ -2357,15 +2361,10 @@ lemma _root_.Set.pairwiseDisjoint_filter [DecidableEq β] (f : α → β) (s : S theorem filter_cons {a : α} (s : Finset α) (ha : a ∉ s) : filter p (cons a s ha) = - (if p a then {a} else ∅ : Finset α).disjUnion (filter p s) - (by - split_ifs - · rw [disjoint_singleton_left] - exact mem_filter.not.mpr <| mt And.left ha - · exact disjoint_empty_left _) := by + if p a then cons a (filter p s) ((mem_of_mem_filter _).mt ha) else filter p s := by split_ifs with h - · rw [filter_cons_of_pos _ _ _ ha h, singleton_disjUnion] - · rw [filter_cons_of_neg _ _ _ ha h, empty_disjUnion] + · rw [filter_cons_of_pos _ _ _ ha h] + · rw [filter_cons_of_neg _ _ _ ha h] section variable [DecidableEq α] @@ -2418,9 +2417,6 @@ lemma filter_and_not (s : Finset α) (p q : α → Prop) [DecidablePred p] [Deci theorem sdiff_eq_filter (s₁ s₂ : Finset α) : s₁ \ s₂ = filter (· ∉ s₂) s₁ := ext fun _ => by simp [mem_sdiff, mem_filter] -theorem sdiff_eq_self (s₁ s₂ : Finset α) : s₁ \ s₂ = s₁ ↔ s₁ ∩ s₂ ⊆ ∅ := by - simp [Subset.antisymm_iff, disjoint_iff_inter_eq_empty] - theorem subset_union_elim {s : Finset α} {t₁ t₂ : Set α} (h : ↑s ⊆ t₁ ∪ t₂) : ∃ s₁ s₂ : Finset α, s₁ ∪ s₂ = s ∧ ↑s₁ ⊆ t₁ ∧ ↑s₂ ⊆ t₂ \ t₁ := by classical @@ -2489,10 +2485,6 @@ theorem filter_ne [DecidableEq β] (s : Finset β) (b : β) : theorem filter_ne' [DecidableEq β] (s : Finset β) (b : β) : (s.filter fun a => a ≠ b) = s.erase b := _root_.trans (filter_congr fun _ _ => by simp_rw [@ne_comm _ b]) (filter_ne s b) -theorem filter_inter_filter_neg_eq (s t : Finset α) : - (s.filter p ∩ t.filter fun a => ¬p a) = ∅ := by - simpa using (disjoint_filter_filter_neg s t p).eq_bot - theorem filter_union_filter_of_codisjoint (s : Finset α) (h : Codisjoint p q) : s.filter p ∪ s.filter q = s := (filter_or _ _ _).symm.trans <| filter_true_of_mem fun x _ => h.top_le x trivial diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 79f88ae3dc5ed..9c66b04ae7c46 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -52,15 +52,14 @@ theorem Iio_castSucc (i : Fin n) : Iio (castSucc i) = (Iio i).map Fin.castSuccEm rw [Finset.map_map, Fin.map_valEmbedding_Iio] exact (Fin.map_valEmbedding_Iio i).symm -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 - rw [Fin.univ_succ, filter_cons, card_disjUnion, filter_map, card_map] - split_ifs <;> simp - 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 := - (card_filter_univ_succ' p).trans (by split_ifs <;> simp [add_comm 1]) + if p 0 then (univ.filter (p ∘ Fin.succ)).card + 1 else (univ.filter (p ∘ Fin.succ)).card := 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 + 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 From 99b8e711e57b074c6653c770dafa640f3be44f34 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 4 Oct 2024 04:46:19 +0000 Subject: [PATCH 167/174] feat(Data/*/Sort): lemmas on sorted lists (#16078) The main thing we prove is that if `l` is a list, and `a` is smaller than all elements on it, then `sort r (a :: l) = a :: sort r l`. --- Mathlib/Data/Finset/Basic.lean | 5 +++++ Mathlib/Data/Finset/Sort.lean | 20 ++++++++++++++++++++ Mathlib/Data/List/Sort.lean | 15 +++++++++++++++ Mathlib/Data/Multiset/Sort.lean | 5 +++++ 4 files changed, 45 insertions(+) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 758307ad5bb02..3e1b0ec9fd4a7 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2912,6 +2912,11 @@ theorem toList_toFinset [DecidableEq α] (s : Finset α) : s.toList.toFinset = s ext simp +theorem _root_.List.toFinset_toList [DecidableEq α] {s : List α} (hs : s.Nodup) : + s.toFinset.toList.Perm s := by + apply List.perm_of_nodup_nodup_toFinset_eq (nodup_toList _) hs + rw [toList_toFinset] + @[simp] theorem toList_eq_singleton_iff {a : α} {s : Finset α} : s.toList = [a] ↔ s = {a} := by rw [toList, Multiset.toList_eq_singleton_iff, val_eq_singleton_iff] diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index dd1aab18ec416..361e4f3721a7f 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -32,6 +32,10 @@ variable (r : α → α → Prop) [DecidableRel r] [IsTrans α r] [IsAntisymm α def sort (s : Finset α) : List α := Multiset.sort r s.1 +@[simp] +theorem sort_val (s : Finset α) : Multiset.sort r s.val = sort r s := + rfl + @[simp] theorem sort_sorted (s : Finset α) : List.Sorted r (sort r s) := Multiset.sort_sorted _ _ @@ -64,11 +68,27 @@ theorem sort_empty : sort r ∅ = [] := theorem sort_singleton (a : α) : sort r {a} = [a] := Multiset.sort_singleton r a +theorem sort_cons {a : α} {s : Finset α} (h₁ : ∀ b ∈ s, r a b) (h₂ : a ∉ s) : + sort r (cons a s h₂) = a :: sort r s := by + rw [sort, cons_val, Multiset.sort_cons r a _ h₁, sort_val] + +theorem sort_insert [DecidableEq α] {a : α} {s : Finset α} (h₁ : ∀ b ∈ s, r a b) (h₂ : a ∉ s) : + sort r (insert a s) = a :: sort r s := by + rw [← cons_eq_insert _ _ h₂, sort_cons r h₁] + open scoped List in theorem sort_perm_toList (s : Finset α) : sort r s ~ s.toList := by rw [← Multiset.coe_eq_coe] simp only [coe_toList, sort_eq] +theorem _root_.List.toFinset_sort [DecidableEq α] {l : List α} (hl : l.Nodup) : + sort r l.toFinset = l ↔ l.Sorted r := by + refine ⟨?_, List.eq_of_perm_of_sorted ((sort_perm_toList r _).trans (List.toFinset_toList hl)) + (sort_sorted r _)⟩ + intro h + rw [← h] + exact sort_sorted r _ + end sort section SortLinearOrder diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 162ca0b8a4138..be93169d29b88 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -198,6 +198,10 @@ def orderedInsert (a : α) : List α → List α | [] => [a] | b :: l => if a ≼ b then a :: b :: l else b :: orderedInsert a l +theorem orderedInsert_of_le {a b : α} (l : List α) (h : a ≼ b) : + orderedInsert r a (b :: l) = a :: b :: l := + dif_pos h + /-- `insertionSort l` returns `l` sorted using the insertion sort algorithm. -/ @[simp] def insertionSort : List α → List α @@ -281,6 +285,17 @@ theorem mem_insertionSort {l : List α} {x : α} : x ∈ l.insertionSort r ↔ x theorem length_insertionSort (l : List α) : (insertionSort r l).length = l.length := (perm_insertionSort r _).length_eq +theorem insertionSort_cons {a : α} {l : List α} (h : ∀ b ∈ l, r a b) : + insertionSort r (a :: l) = a :: insertionSort r l := by + rw [insertionSort] + cases hi : insertionSort r l with + | nil => rfl + | cons b m => + rw [orderedInsert_of_le] + apply h b <| (mem_insertionSort r).1 _ + rw [hi] + exact mem_cons_self b m + theorem map_insertionSort (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b ∈ l, a ≼ b ↔ f a ≼ f b) : (l.insertionSort r).map f = (l.map f).insertionSort s := by induction l with diff --git a/Mathlib/Data/Multiset/Sort.lean b/Mathlib/Data/Multiset/Sort.lean index 545d09d151580..997af93b548c0 100644 --- a/Mathlib/Data/Multiset/Sort.lean +++ b/Mathlib/Data/Multiset/Sort.lean @@ -62,6 +62,11 @@ theorem map_sort (f : α → β) (s : Multiset α) revert s exact Quot.ind fun _ => List.map_mergeSort' _ _ _ _ +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 + end sort -- TODO: use a sort order if available, gh-18166 From 7e65505042f561db2a6fe1c24a09d1d2c5a0df5b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 4 Oct 2024 04:46:20 +0000 Subject: [PATCH 168/174] feat: Order instances for `MulOpposite`/`AddOpposite` (#17201) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Transfer order and ordered monoid/group/semiring/ring instances from `α` to `αᵐᵒᵖ` and `αᵃᵒᵖ`. --- Mathlib.lean | 2 + Mathlib/Algebra/Order/Group/Opposite.lean | 85 +++++++++++++++++++++++ Mathlib/Algebra/Order/Ring/Opposite.lean | 50 +++++++++++++ 3 files changed, 137 insertions(+) create mode 100644 Mathlib/Algebra/Order/Group/Opposite.lean create mode 100644 Mathlib/Algebra/Order/Ring/Opposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 58afee637aa81..098d26089a621 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -593,6 +593,7 @@ import Mathlib.Algebra.Order.Group.Int import Mathlib.Algebra.Order.Group.Lattice import Mathlib.Algebra.Order.Group.MinMax import Mathlib.Algebra.Order.Group.Nat +import Mathlib.Algebra.Order.Group.Opposite import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Algebra.Order.Group.PiLex import Mathlib.Algebra.Order.Group.Pointwise.Bounds @@ -667,6 +668,7 @@ import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.Order.Ring.Opposite import Mathlib.Algebra.Order.Ring.Pow import Mathlib.Algebra.Order.Ring.Prod import Mathlib.Algebra.Order.Ring.Rat diff --git a/Mathlib/Algebra/Order/Group/Opposite.lean b/Mathlib/Algebra/Order/Group/Opposite.lean new file mode 100644 index 0000000000000..e145f616c4c11 --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Opposite.lean @@ -0,0 +1,85 @@ +/- +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.Order.Group.Defs +import Mathlib.Algebra.Group.Opposite + +/-! +# Order instances for `MulOpposite`/`AddOpposite` + +This files transfers order instances and ordered monoid/group instances from `α` to `αᵐᵒᵖ` and +`αᵃᵒᵖ`. +-/ + +variable {α : Type*} + +namespace MulOpposite +section Preorder +variable [Preorder α] + +@[to_additive] instance : Preorder αᵐᵒᵖ := Preorder.lift unop + +@[to_additive (attr := simp)] lemma unop_le_unop {a b : αᵐᵒᵖ} : a.unop ≤ b.unop ↔ a ≤ b := .rfl +@[to_additive (attr := simp)] lemma op_le_op {a b : α} : op a ≤ op b ↔ a ≤ b := .rfl + +end Preorder + +@[to_additive] instance [PartialOrder α] : PartialOrder αᵐᵒᵖ := PartialOrder.lift _ unop_injective + +section OrderedCommMonoid +variable [OrderedCommMonoid α] + +@[to_additive] instance : OrderedCommMonoid αᵐᵒᵖ where + mul_le_mul_left a b hab c := mul_le_mul_right' (by simpa) c.unop + +@[to_additive (attr := simp)] lemma unop_le_one {a : αᵐᵒᵖ} : unop a ≤ 1 ↔ a ≤ 1 := .rfl +@[to_additive (attr := simp)] lemma one_le_unop {a : αᵐᵒᵖ} : 1 ≤ unop a ↔ 1 ≤ a := .rfl +@[to_additive (attr := simp)] lemma op_le_one {a : α} : op a ≤ 1 ↔ a ≤ 1 := .rfl +@[to_additive (attr := simp)] lemma one_le_op {a : α} : 1 ≤ op a ↔ 1 ≤ a := .rfl + +end OrderedCommMonoid + +@[to_additive] instance [OrderedCommGroup α] : OrderedCommGroup αᵐᵒᵖ where + __ := instCommGroup + __ := instOrderedCommMonoid + +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid α] + +instance : OrderedAddCommMonoid αᵐᵒᵖ where + add_le_add_left a b hab c := add_le_add_left (by simpa) c.unop + +@[simp] lemma unop_nonneg {a : αᵐᵒᵖ} : unop a ≤ 0 ↔ a ≤ 0 := .rfl +@[simp] lemma unop_nonpos {a : αᵐᵒᵖ} : 0 ≤ unop a ↔ 0 ≤ a := .rfl +@[simp] lemma op_nonneg {a : α} : op a ≤ 0 ↔ a ≤ 0 := .rfl +@[simp] lemma op_nonpos {a : α} : 0 ≤ op a ↔ 0 ≤ a := .rfl + +end OrderedAddCommMonoid + +instance [OrderedAddCommGroup α] : OrderedAddCommGroup αᵐᵒᵖ where + __ := instAddCommGroup + __ := instOrderedAddCommMonoid + +end MulOpposite + +namespace AddOpposite +section OrderedCommMonoid +variable [OrderedCommMonoid α] + +instance : OrderedCommMonoid αᵃᵒᵖ where + mul_le_mul_left a b hab c := mul_le_mul_left' (by simpa) c.unop + +@[simp] lemma unop_le_one {a : αᵃᵒᵖ} : unop a ≤ 1 ↔ a ≤ 1 := .rfl +@[simp] lemma one_le_unop {a : αᵃᵒᵖ} : 1 ≤ unop a ↔ 1 ≤ a := .rfl +@[simp] lemma op_le_one {a : α} : op a ≤ 1 ↔ a ≤ 1 := .rfl +@[simp] lemma one_le_op {a : α} : 1 ≤ op a ↔ 1 ≤ a := .rfl + +end OrderedCommMonoid + +instance [OrderedCommGroup α] : OrderedCommGroup αᵃᵒᵖ where + __ := instCommGroup + __ := instOrderedCommMonoid + +end AddOpposite diff --git a/Mathlib/Algebra/Order/Ring/Opposite.lean b/Mathlib/Algebra/Order/Ring/Opposite.lean new file mode 100644 index 0000000000000..c945c5ec82242 --- /dev/null +++ b/Mathlib/Algebra/Order/Ring/Opposite.lean @@ -0,0 +1,50 @@ +/- +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.Order.Group.Opposite +import Mathlib.Algebra.Order.Ring.Defs +import Mathlib.Algebra.Ring.Opposite + +/-! +# Ordered ring instances for `MulOpposite`/`AddOpposite` + +This files transfers ordered (semi)ring instances from `α` to `αᵐᵒᵖ` and `αᵃᵒᵖ`. +-/ + +variable {α : Type*} + +namespace MulOpposite + +instance [OrderedSemiring α] : OrderedSemiring αᵐᵒᵖ where + __ := instSemiring + __ := instOrderedAddCommMonoid + zero_le_one := zero_le_one (α := α) + mul_le_mul_of_nonneg_left _ _ _ := mul_le_mul_of_nonneg_right (α := α) + mul_le_mul_of_nonneg_right _ _ _ := mul_le_mul_of_nonneg_left (α := α) + +instance [OrderedRing α] : OrderedRing αᵐᵒᵖ where + __ := instRing + __ := instOrderedAddCommGroup + __ := instOrderedSemiring + mul_nonneg _a _b ha hb := mul_nonneg (α := α) hb ha + +end MulOpposite + +namespace AddOpposite + +instance [OrderedSemiring α] : OrderedSemiring αᵃᵒᵖ where + __ := instSemiring + __ := instOrderedAddCommMonoid + zero_le_one := zero_le_one (α := α) + mul_le_mul_of_nonneg_left _ _ _ := mul_le_mul_of_nonneg_left (α := α) + mul_le_mul_of_nonneg_right _ _ _ := mul_le_mul_of_nonneg_right (α := α) + +instance [OrderedRing α] : OrderedRing αᵐᵒᵖ where + __ := instRing + __ := instOrderedAddCommGroup + __ := instOrderedSemiring + mul_nonneg _a _b := mul_nonneg (α := α) + +end AddOpposite From 6dd0053d1d38ec7c6c0dd7bb5b5f72ab271c005e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 4 Oct 2024 04:46:22 +0000 Subject: [PATCH 169/174] feat: Big operators indexed by an interval (#17280) From LeanCamCombi --- Mathlib.lean | 1 + .../BigOperators/Group/LocallyFinite.lean | 80 +++++++++++++++++++ 2 files changed, 81 insertions(+) create mode 100644 Mathlib/Algebra/Order/BigOperators/Group/LocallyFinite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 098d26089a621..79c393e06f29b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -552,6 +552,7 @@ import Mathlib.Algebra.Order.Archimedean.Submonoid import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.List +import Mathlib.Algebra.Order.BigOperators.Group.LocallyFinite import Mathlib.Algebra.Order.BigOperators.Group.Multiset import Mathlib.Algebra.Order.BigOperators.GroupWithZero.List import Mathlib.Algebra.Order.BigOperators.GroupWithZero.Multiset diff --git a/Mathlib/Algebra/Order/BigOperators/Group/LocallyFinite.lean b/Mathlib/Algebra/Order/BigOperators/Group/LocallyFinite.lean new file mode 100644 index 0000000000000..f43f25244887d --- /dev/null +++ b/Mathlib/Algebra/Order/BigOperators/Group/LocallyFinite.lean @@ -0,0 +1,80 @@ +/- +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.BigOperators.Group.Finset +import Mathlib.Order.Interval.Finset.Basic + +/-! +# Big operators indexed by intervals + +This file proves lemmas about `∏ x ∈ Ixx a b, f x` and `∑ x ∈ Ixx a b, f x`. +-/ + +variable {α β : Type*} [PartialOrder α] [CommMonoid β] {f : α → β} {a b : α} + +namespace Finset +section LocallyFiniteOrder +variable [LocallyFiniteOrder α] + +@[to_additive] +lemma left_mul_prod_Ioc (h : a ≤ b) : f a * ∏ x ∈ Ioc a b, f x = ∏ x ∈ Icc a b, f x := by + rw [Icc_eq_cons_Ioc h, prod_cons] + +@[to_additive] +lemma prod_Ioc_mul_left (h : a ≤ b) : (∏ x ∈ Ioc a b, f x) * f a = ∏ x ∈ Icc a b, f x := by + rw [mul_comm, left_mul_prod_Ioc h] + +@[to_additive] +lemma right_mul_prod_Ico (h : a ≤ b) : f b * ∏ x ∈ Ico a b, f x = ∏ x ∈ Icc a b, f x := by + rw [Icc_eq_cons_Ico h, prod_cons] + +@[to_additive] +lemma prod_Ico_mul_right (h : a ≤ b) : (∏ x ∈ Ico a b, f x) * f b = ∏ x ∈ Icc a b, f x := by + rw [mul_comm, right_mul_prod_Ico h] + +@[to_additive] +lemma left_mul_prod_Ioo (h : a < b) : f a * ∏ x ∈ Ioo a b, f x = ∏ x ∈ Ico a b, f x := by + rw [Ico_eq_cons_Ioo h, prod_cons] + +@[to_additive] +lemma prod_Ioo_mul_left (h : a < b) : (∏ x ∈ Ioo a b, f x) * f a = ∏ x ∈ Ico a b, f x := by + rw [mul_comm, left_mul_prod_Ioo h] + +@[to_additive] +lemma right_mul_prod_Ioo (h : a < b) : f b * ∏ x ∈ Ioo a b, f x = ∏ x ∈ Ioc a b, f x := by + rw [Ioc_eq_cons_Ioo h, prod_cons] + +@[to_additive] +lemma prod_Ioo_mul_right (h : a < b) : (∏ x ∈ Ioo a b, f x) * f b = ∏ x ∈ Ioc a b, f x := by + rw [mul_comm, right_mul_prod_Ioo h] + +end LocallyFiniteOrder + +section LocallyFiniteOrderTop +variable [LocallyFiniteOrderTop α] + +@[to_additive] +lemma left_mul_prod_Ioi (a : α) : f a * ∏ x ∈ Ioi a, f x = ∏ x ∈ Ici a, f x := by + rw [Ici_eq_cons_Ioi, prod_cons] + +@[to_additive] +lemma prod_Ioi_mul_left (a : α) : (∏ x ∈ Ioi a, f x) * f a = ∏ x ∈ Ici a, f x := by + rw [mul_comm, left_mul_prod_Ioi] + +end LocallyFiniteOrderTop + +section LocallyFiniteOrderBot +variable [LocallyFiniteOrderBot α] + +@[to_additive] +lemma right_mul_prod_Iio (a : α) : f a * ∏ x ∈ Iio a, f x = ∏ x ∈ Iic a, f x := by + rw [Iic_eq_cons_Iio, prod_cons] + +@[to_additive] +lemma prod_Iio_mul_right (a : α) : (∏ x ∈ Iio a, f x) * f a = ∏ x ∈ Iic a, f x := by + rw [mul_comm, right_mul_prod_Iio] + +end LocallyFiniteOrderBot +end Finset From 149d450dab426eecce5b88614156d8009022c86e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 4 Oct 2024 04:46:23 +0000 Subject: [PATCH 170/174] feat: Order properties of `Pi.single` (#17281) From LeanCamCombi --- Mathlib/Algebra/Order/Pi.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/Algebra/Order/Pi.lean b/Mathlib/Algebra/Order/Pi.lean index be50955c5d5ae..9710857e2382b 100644 --- a/Mathlib/Algebra/Order/Pi.lean +++ b/Mathlib/Algebra/Order/Pi.lean @@ -128,6 +128,25 @@ variable [One γ] [LE γ] {f : α → β} {g : α → γ} {e : β → γ} end extend end Function + +namespace Pi +variable {ι : Type*} {α : ι → Type*} [DecidableEq ι] [∀ i, One (α i)] [∀ i, Preorder (α i)] {i : ι} + {a b : α i} + +@[to_additive (attr := simp)] +lemma mulSingle_le_mulSingle : mulSingle i a ≤ mulSingle i b ↔ a ≤ b := by + simp [mulSingle, update_le_update_iff] + +@[to_additive (attr := gcongr)] alias ⟨_, GCongr.mulSingle_mono⟩ := mulSingle_le_mulSingle + +@[to_additive (attr := simp) single_nonneg] +lemma one_le_mulSingle : 1 ≤ mulSingle i a ↔ 1 ≤ a := by simp [mulSingle] + +@[to_additive (attr := simp)] +lemma mulSingle_le_one : mulSingle i a ≤ 1 ↔ a ≤ 1 := by simp [mulSingle] + +end Pi + -- Porting note: Tactic code not ported yet -- namespace Tactic From a083bae251dee28c8a1efbb012b21c77c388f63e Mon Sep 17 00:00:00 2001 From: Wojciech Nawrocki Date: Fri, 4 Oct 2024 04:46:24 +0000 Subject: [PATCH 171/174] feat(RingTheory): define two-sided Jacobson radical (#17341) - Add a constructor `TwoSidedIdeal.ofIdeal` to turn an `Ideal` that happens to be two-sided into a `TwoSidedIdeal`. - Use `ofIdeal` to define a two-sided version of the Jacobson radical. - Generalize some ring lemmas that didn't need to assume commutativity. Co-authored-by: Wojciech Nawrocki --- Mathlib/RingTheory/Ideal/Basic.lean | 10 +-- Mathlib/RingTheory/JacobsonIdeal.lean | 78 +++++++++++++++++-- Mathlib/RingTheory/TwoSidedIdeal/Basic.lean | 15 ++-- .../RingTheory/TwoSidedIdeal/Operations.lean | 41 ++++++++-- 4 files changed, 117 insertions(+), 27 deletions(-) diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 3b495e29ee8df..36081a3d2b091 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -149,6 +149,9 @@ theorem mem_span_insert {s : Set α} {x y} : theorem mem_span_singleton' {x y : α} : x ∈ span ({y} : Set α) ↔ ∃ a, a * y = x := Submodule.mem_span_singleton +theorem mem_span_singleton_self (x : α) : x ∈ span ({x} : Set α) := + Submodule.mem_span_singleton_self x + theorem span_singleton_le_iff_mem {x : α} : span {x} ≤ I ↔ x ∈ I := Submodule.span_singleton_le_iff_mem _ _ @@ -184,8 +187,8 @@ theorem span_eq_top_iff_finite (s : Set α) : simp_rw [eq_top_iff_one] exact ⟨Submodule.mem_span_finite_of_mem_span, fun ⟨s', h₁, h₂⟩ => span_mono h₁ h₂⟩ -theorem mem_span_singleton_sup {S : Type*} [CommSemiring S] {x y : S} {I : Ideal S} : - x ∈ Ideal.span {y} ⊔ I ↔ ∃ a : S, ∃ b ∈ I, a * y + b = x := by +theorem mem_span_singleton_sup {x y : α} {I : Ideal α} : + x ∈ Ideal.span {y} ⊔ I ↔ ∃ a : α, ∃ b ∈ I, a * y + b = x := by rw [Submodule.mem_sup] constructor · rintro ⟨ya, hya, b, hb, rfl⟩ @@ -431,9 +434,6 @@ theorem mul_unit_mem_iff_mem {x y : α} (hy : IsUnit y) : x * y ∈ I ↔ x ∈ theorem mem_span_singleton {x y : α} : x ∈ span ({y} : Set α) ↔ y ∣ x := mem_span_singleton'.trans <| exists_congr fun _ => by rw [eq_comm, mul_comm] -theorem mem_span_singleton_self (x : α) : x ∈ span ({x} : Set α) := - mem_span_singleton.mpr dvd_rfl - theorem span_singleton_le_span_singleton {x y : α} : span ({x} : Set α) ≤ span ({y} : Set α) ↔ y ∣ x := span_le.trans <| singleton_subset_iff.trans mem_span_singleton diff --git a/Mathlib/RingTheory/JacobsonIdeal.lean b/Mathlib/RingTheory/JacobsonIdeal.lean index c51eef801b35e..5f513287fa395 100644 --- a/Mathlib/RingTheory/JacobsonIdeal.lean +++ b/Mathlib/RingTheory/JacobsonIdeal.lean @@ -1,11 +1,12 @@ /- Copyright (c) 2020 Devon Tuma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau, Devon Tuma +Authors: Kenny Lau, Devon Tuma, Wojciech Nawrocki -/ import Mathlib.RingTheory.Ideal.IsPrimary import Mathlib.RingTheory.Ideal.Quotient import Mathlib.RingTheory.Polynomial.Quotient +import Mathlib.RingTheory.TwoSidedIdeal.Operations /-! # Jacobson radical @@ -13,20 +14,24 @@ import Mathlib.RingTheory.Polynomial.Quotient The Jacobson radical of a ring `R` is defined to be the intersection of all maximal ideals of `R`. This is similar to how the nilradical is equal to the intersection of all prime ideals of `R`. -We can extend the idea of the nilradical to ideals of `R`, -by letting the radical of an ideal `I` be the intersection of prime ideals containing `I`. +We can extend the idea of the nilradical of `R` to ideals of `R`, +by letting the nilradical of an ideal `I` be the intersection of prime ideals containing `I`. Under this extension, the original nilradical is the radical of the zero ideal `⊥`. Here we define the Jacobson radical of an ideal `I` in a similar way, as the intersection of maximal ideals containing `I`. ## Main definitions -Let `R` be a commutative ring, and `I` be an ideal of `R` +Let `R` be a ring, and `I` be a left ideal of `R` -* `Ideal.jacobson I` is the jacobson radical, i.e. the infimum of all maximal ideals containing I. +* `Ideal.jacobson I` is the Jacobson radical, i.e. the infimum of all maximal ideals containing `I`. * `Ideal.IsLocal I` is the proposition that the jacobson radical of `I` is itself a maximal ideal +Furthermore when `I` is a two-sided ideal of `R` + +* `TwoSidedIdeal.jacobson I` is the Jacobson radical as a two-sided ideal + ## Main statements * `mem_jacobson_iff` gives a characterization of members of the jacobson of I @@ -111,11 +116,17 @@ theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x sub_add_cancel] exact M.mul_mem_left _ hi) <| him hz⟩ -theorem exists_mul_sub_mem_of_sub_one_mem_jacobson {I : Ideal R} (r : R) (h : r - 1 ∈ jacobson I) : - ∃ s, s * r - 1 ∈ I := by +theorem exists_mul_add_sub_mem_of_mem_jacobson {I : Ideal R} (r : R) (h : r ∈ jacobson I) : + ∃ s, s * (r + 1) - 1 ∈ I := by cases' mem_jacobson_iff.1 h 1 with s hs use s - simpa [mul_sub] using hs + rw [mul_add, mul_one] + simpa using hs + +theorem exists_mul_sub_mem_of_sub_one_mem_jacobson {I : Ideal R} (r : R) (h : r - 1 ∈ jacobson I) : + ∃ s, s * r - 1 ∈ I := by + convert exists_mul_add_sub_mem_of_mem_jacobson _ h + simp /-- An ideal equals its Jacobson radical iff it is the intersection of a set of maximal ideals. Allowing the set to include ⊤ is equivalent, and is included only to simplify some proofs. -/ @@ -214,6 +225,44 @@ theorem jacobson_mono {I J : Ideal R} : I ≤ J → I.jacobson ≤ J.jacobson := erw [mem_sInf] at hx ⊢ exact fun K ⟨hK, hK_max⟩ => hx ⟨Trans.trans h hK, hK_max⟩ +/-- The Jacobson radical of a two-sided ideal is two-sided. + +It is preferable to use `TwoSidedIdeal.jacobson` instead of this lemma. -/ +theorem jacobson_mul_mem_right {I : Ideal R} + (mul_mem_right : ∀ {x y}, x ∈ I → x * y ∈ I) : + ∀ {x y}, x ∈ I.jacobson → x * y ∈ I.jacobson := by + -- Proof generalized from + -- https://ysharifi.wordpress.com/2022/08/16/the-jacobson-radical-definition-and-basic-results/ + intro x r xJ + apply mem_sInf.mpr + intro 𝔪 𝔪_mem + by_cases r𝔪 : r ∈ 𝔪 + · apply 𝔪.smul_mem _ r𝔪 + -- 𝔪₀ := { a : R | a*r ∈ 𝔪 } + let 𝔪₀ : Ideal R := Submodule.comap (DistribMulAction.toLinearMap R (S := Rᵐᵒᵖ) R (.op r)) 𝔪 + suffices x ∈ 𝔪₀ by simpa [𝔪₀] using this + have I𝔪₀ : I ≤ 𝔪₀ := fun i iI => + 𝔪_mem.left (mul_mem_right iI) + have 𝔪₀_maximal : IsMaximal 𝔪₀ := by + refine isMaximal_iff.mpr ⟨ + fun h => r𝔪 (by simpa [𝔪₀] using h), + fun J b 𝔪₀J b𝔪₀ bJ => ?_⟩ + let K : Ideal R := Ideal.span {b*r} ⊔ 𝔪 + have ⟨s, y, y𝔪, sbyr⟩ := + mem_span_singleton_sup.mp <| + mul_mem_left _ r <| + (isMaximal_iff.mp 𝔪_mem.right).right K (b*r) + le_sup_right b𝔪₀ + (mem_sup_left <| mem_span_singleton_self _) + have : 1 - s*b ∈ 𝔪₀ := by + rw [mul_one, add_comm, ← eq_sub_iff_add_eq] at sbyr + rw [sbyr, ← mul_assoc] at y𝔪 + simp [𝔪₀, sub_mul, y𝔪] + have : 1 - s*b + s*b ∈ J := by + apply add_mem (𝔪₀J this) (J.mul_mem_left _ bJ) + simpa using this + exact mem_sInf.mp xJ ⟨I𝔪₀, 𝔪₀_maximal⟩ + end Ring section CommRing @@ -375,3 +424,16 @@ theorem isPrimary_of_isMaximal_radical [CommRing R] {I : Ideal R} (hi : IsMaxima (this ▸ id)⟩ end Ideal + +namespace TwoSidedIdeal + +variable {R : Type u} [Ring R] + +/-- The Jacobson radical of `I` is the infimum of all maximal (left) ideals containing `I`. -/ +def jacobson (I : TwoSidedIdeal R) : TwoSidedIdeal R := + (asIdeal I).jacobson.toTwoSided (Ideal.jacobson_mul_mem_right <| I.mul_mem_right _ _) + +lemma asIdeal_jacobson (I : TwoSidedIdeal R) : asIdeal I.jacobson = (asIdeal I).jacobson := by + ext; simp [jacobson] + +end TwoSidedIdeal diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean index 8be0b86366462..dd7cee192929b 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean @@ -156,17 +156,18 @@ def mk' (carrier : Set R) rw [show a + c - (b + d) = (a - b) + (c - d) by abel] exact add_mem h1 h2 } -lemma mem_mk' (carrier : Set R) - (zero_mem : 0 ∈ carrier) - (add_mem : ∀ {x y}, x ∈ carrier → y ∈ carrier → x + y ∈ carrier) - (neg_mem : ∀ {x}, x ∈ carrier → -x ∈ carrier) - (mul_mem_left : ∀ {x y}, y ∈ carrier → x * y ∈ carrier) - (mul_mem_right : ∀ {x y}, x ∈ carrier → x * y ∈ carrier) - (x : R) : +@[simp] +lemma mem_mk' (carrier : Set R) (zero_mem add_mem neg_mem mul_mem_left mul_mem_right) (x : R) : x ∈ mk' carrier zero_mem add_mem neg_mem mul_mem_left mul_mem_right ↔ x ∈ carrier := by rw [mem_iff] simp [mk'] +set_option linter.docPrime false in +@[simp] +lemma coe_mk' (carrier : Set R) (zero_mem add_mem neg_mem mul_mem_left mul_mem_right) : + (mk' carrier zero_mem add_mem neg_mem mul_mem_left mul_mem_right : Set R) = carrier := + Set.ext <| mem_mk' carrier zero_mem add_mem neg_mem mul_mem_left mul_mem_right + instance : SMulMemClass (TwoSidedIdeal R) R R where smul_mem _ _ h := TwoSidedIdeal.mul_mem_left _ _ _ h diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean index 455172552a77c..78ccd01d0a217 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean @@ -106,12 +106,7 @@ def ker : TwoSidedIdeal R := (by rintro _ _ (h : f _ = 0); simp [h]) lemma mem_ker {x : R} : x ∈ ker f ↔ f x = 0 := by - delta ker; rw [mem_mk'] - · rfl - · rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2] - · rintro _ (h : f _ = 0); simp [h] - · rintro _ _ (h : f _ = 0); simp [h] - · rintro _ _ (h : f _ = 0); simp [h] + delta ker; rw [mem_mk']; rfl end NonUnitalNonAssocRing @@ -226,6 +221,7 @@ def asIdeal : TwoSidedIdeal R →o Ideal R where smul_mem' := fun r x hx => I.mul_mem_left r x hx } monotone' _ _ h _ h' := h h' +@[simp] lemma mem_asIdeal {I : TwoSidedIdeal R} {x : R} : x ∈ asIdeal I ↔ x ∈ I := by simp [asIdeal] @@ -266,8 +262,39 @@ def orderIsoIdeal : TwoSidedIdeal R ≃o Ideal R where right_inv J := SetLike.ext fun x ↦ mem_span_iff.trans ⟨fun h ↦ mem_mk' _ _ _ _ _ _ _ |>.1 <| h (mk' J J.zero_mem J.add_mem J.neg_mem (J.mul_mem_left _) (J.mul_mem_right _)) - (fun x => by simp [mem_mk']), by aesop⟩ + (fun x => by simp), by aesop⟩ end CommRing end TwoSidedIdeal + +namespace Ideal +variable {R : Type*} [Ring R] + +/-- Bundle an `Ideal` that is already two-sided as a `TwoSidedIdeal`. -/ +def toTwoSided (I : Ideal R) (mul_mem_right : ∀ {x y}, x ∈ I → x * y ∈ I) : TwoSidedIdeal R := + TwoSidedIdeal.mk' I I.zero_mem I.add_mem I.neg_mem (I.smul_mem _) mul_mem_right + +@[simp] +lemma mem_toTwoSided {I : Ideal R} {h} {x : R} : + x ∈ I.toTwoSided h ↔ x ∈ I := by + simp [toTwoSided] + +@[simp] +lemma coe_toTwoSided (I : Ideal R) (h) : (I.toTwoSided h : Set R) = I := by + simp [toTwoSided] + +@[simp] +lemma toTwoSided_asIdeal (I : TwoSidedIdeal R) (h) : (TwoSidedIdeal.asIdeal I).toTwoSided h = I := + by ext; simp + +@[simp] +lemma asIdeal_toTwoSided (I : Ideal R) (h) : TwoSidedIdeal.asIdeal (I.toTwoSided h) = I := by + ext + simp + +instance : CanLift (Ideal R) (TwoSidedIdeal R) TwoSidedIdeal.asIdeal + (fun I => ∀ {x y}, x ∈ I → x * y ∈ I) where + prf I mul_mem_right := ⟨I.toTwoSided mul_mem_right, asIdeal_toTwoSided ..⟩ + +end Ideal From 630c632468b780bb80ee0f04514ef67b1167412b Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Fri, 4 Oct 2024 04:46:25 +0000 Subject: [PATCH 172/174] refactor(Algebra/Field/Subfield): minor golfing (#17376) --- Mathlib/Algebra/Field/Subfield.lean | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/Field/Subfield.lean b/Mathlib/Algebra/Field/Subfield.lean index 63adf4ac50a55..30b6731d7f52c 100644 --- a/Mathlib/Algebra/Field/Subfield.lean +++ b/Mathlib/Algebra/Field/Subfield.lean @@ -126,20 +126,20 @@ variable (S) /-- A subfield inherits a division ring structure -/ instance (priority := 75) toDivisionRing (s : S) : DivisionRing s := Subtype.coe_injective.divisionRing ((↑) : s → K) - (by rfl) (by rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) - (by intros; rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) - (by intros; rfl) (coe_nnqsmul _) (coe_qsmul _) (by intros; rfl) (by intros; rfl) - (by intros; rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) + rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) + (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + (fun _ _ ↦ rfl) (coe_nnqsmul _) (coe_qsmul _) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + (fun _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) -- Prefer subclasses of `Field` over subclasses of `SubfieldClass`. /-- A subfield of a field inherits a field structure -/ instance (priority := 75) toField {K} [Field K] [SetLike S K] [SubfieldClass S K] (s : S) : Field s := Subtype.coe_injective.field ((↑) : s → K) - (by rfl) (by rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) - (by intros; rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) (by intros; rfl) - (coe_nnqsmul _) (coe_qsmul _) (by intros; rfl) (by intros; rfl) (by intros; rfl) - (by intros; rfl) (by intros; rfl) (by intros; rfl) + rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) + (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + (coe_nnqsmul _) (coe_qsmul _) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) + (fun _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) end SubfieldClass @@ -313,15 +313,15 @@ instance : Pow s ℤ := instance toDivisionRing (s : Subfield K) : DivisionRing s := Subtype.coe_injective.divisionRing ((↑) : s → K) rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) - (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (by intros; rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) - (by intros; rfl) fun _ ↦ rfl + (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) + (fun _ ↦ rfl) fun _ ↦ rfl /-- A subfield inherits a field structure -/ instance toField {K} [Field K] (s : Subfield K) : Field s := Subtype.coe_injective.field ((↑) : s → K) rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) - (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (by intros; rfl) (fun _ => rfl) - (fun _ => rfl) (by intros; rfl) fun _ => rfl + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ ↦ rfl) (fun _ => rfl) + (fun _ => rfl) (fun _ ↦ rfl) fun _ => rfl @[simp, norm_cast] theorem coe_add (x y : s) : (↑(x + y) : K) = ↑x + ↑y := From b0bcf46fb80c3cb1eba7696aadf4ee19b959fd09 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 4 Oct 2024 06:17:50 +0000 Subject: [PATCH 173/174] feat: if a function is analytic on a set, its derivative also is, even if the space is not complete (#17221) We already have a version of this theorem, but assuming completeness while this is not necessary: if the function is differentiable, then the power series for its derivative converges, to the given differential (since this is the case in the completion, and the embedding in the completion is an embedding). This result requires expanding the API around derivatives of analytic functions. As a byproduct, we also write down the derivative of linear maps into multilinear maps (which will be needed for the Faa di Bruno formula). --- .../Analysis/Calculus/FDeriv/Analytic.lean | 405 +++++++++++++++++- .../Normed/Operator/LinearIsometry.lean | 2 + .../NormedSpace/OperatorNorm/NormedSpace.lean | 9 + 3 files changed, 395 insertions(+), 21 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 92be2e3cc4f08..0e90e3c0a221f 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -4,10 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Analytic.CPolynomial +import Mathlib.Analysis.Analytic.Inverse import Mathlib.Analysis.Analytic.Within import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.Calculus.ContDiff.Defs import Mathlib.Analysis.Calculus.FDeriv.Add +import Mathlib.Analysis.Calculus.FDeriv.Prod +import Mathlib.Analysis.Normed.Module.Completion /-! # Frechet derivatives of analytic functions. @@ -17,11 +20,50 @@ Also the special case in terms of `deriv` when the domain is 1-dimensional. As an application, we show that continuous multilinear maps are smooth. We also compute their iterated derivatives, in `ContinuousMultilinearMap.iteratedFDeriv_eq`. + +## Main definitions and results + +* `AnalyticAt.differentiableAt` : an analytic function at a point is differentiable there. +* `AnalyticOnNhd.fderiv` : in a complete space, if a function is analytic on a + neighborhood of a set `s`, so is its derivative. +* `AnalyticOnNhd.fderiv_of_isOpen` : if a function is analytic on a neighborhood of an + open set `s`, so is its derivative. +* `AnalyticOn.fderivWithin` : if a function is analytic on a set of unique differentiability, + so is its derivative within this set. +* `PartialHomeomorph.analyticAt_symm` : if a partial homeomorphism `f` is analytic at a + point `f.symm a`, with invertible derivative, then its inverse is analytic at `a`. + +## Comments on completeness + +Some theorems need a complete space, some don't, for the following reason. + +(1) If a function is analytic at a point `x`, then it is differentiable there (with derivative given +by the first term in the power series). There is no issue of convergence here. + +(2) If a function has a power series on a ball `B (x, r)`, there is no guarantee that the power +series for the derivative will converge at `y ≠ x`, if the space is not complete. So, to deduce +that `f` is differentiable at `y`, one needs completeness in general. + +(3) However, if a function `f` has a power series on a ball `B (x, r)`, and is a priori known to be +differentiable at some point `y ≠ x`, then the power series for the derivative of `f` will +automatically converge at `y`, towards the given derivative: this follows from the facts that this +is true in the completion (thanks to the previous point) and that the map to the completion is +an embedding. + +(4) Therefore, if one assumes `AnalyticOn 𝕜 f s` where `s` is an open set, then `f` is analytic +therefore differentiable at every point of `s`, by (1), so by (3) the power series for its +derivative converges on whole balls. Therefore, the derivative of `f` is also analytic on `s`. The +same holds if `s` is merely a set with unique differentials. + +(5) However, this does not work for `AnalyticOnNhd 𝕜 f s`, as we don't get for free +differentiability at points in a neighborhood of `s`. Therefore, the theorem that deduces +`AnalyticOnNhd 𝕜 (fderiv 𝕜 f) s` from `AnalyticOnNhd 𝕜 f s` requires completeness of the space. + -/ open Filter Asymptotics Set -open scoped ENNReal +open scoped ENNReal Topology universe u v @@ -34,26 +76,58 @@ section fderiv variable {p : FormalMultilinearSeries 𝕜 E F} {r : ℝ≥0∞} variable {f : E → F} {x : E} {s : Set E} -theorem HasFPowerSeriesAt.hasStrictFDerivAt (h : HasFPowerSeriesAt f p x) : - HasStrictFDerivAt f (continuousMultilinearCurryFin1 𝕜 E F (p 1)) x := by +/-- A function which is analytic within a set is strictly differentiable there. Since we +don't have a predicate `HasStrictFDerivWithinAt`, we spell out what it would mean. -/ +theorem HasFPowerSeriesWithinAt.hasStrictFDerivWithinAt (h : HasFPowerSeriesWithinAt f p s x) : + (fun y ↦ f y.1 - f y.2 - ((continuousMultilinearCurryFin1 𝕜 E F) (p 1)) (y.1 - y.2)) + =o[𝓝[insert x s ×ˢ insert x s] (x, x)] fun y ↦ y.1 - y.2 := by refine h.isBigO_image_sub_norm_mul_norm_sub.trans_isLittleO (IsLittleO.of_norm_right ?_) refine isLittleO_iff_exists_eq_mul.2 ⟨fun y => ‖y - (x, x)‖, ?_, EventuallyEq.rfl⟩ + apply Tendsto.mono_left _ nhdsWithin_le_nhds refine (continuous_id.sub continuous_const).norm.tendsto' _ _ ?_ rw [_root_.id, sub_self, norm_zero] +theorem HasFPowerSeriesAt.hasStrictFDerivAt (h : HasFPowerSeriesAt f p x) : + HasStrictFDerivAt f (continuousMultilinearCurryFin1 𝕜 E F (p 1)) x := by + simpa only [Set.insert_eq_of_mem, Set.mem_univ, Set.univ_prod_univ, nhdsWithin_univ] + using (h.hasFPowerSeriesWithinAt (s := Set.univ)).hasStrictFDerivWithinAt + +theorem HasFPowerSeriesWithinAt.hasFDerivWithinAt (h : HasFPowerSeriesWithinAt f p s x) : + HasFDerivWithinAt f (continuousMultilinearCurryFin1 𝕜 E F (p 1)) (insert x s) x := by + rw [HasFDerivWithinAt, hasFDerivAtFilter_iff_isLittleO, isLittleO_iff] + intro c hc + have : Tendsto (fun y ↦ (y, x)) (𝓝[insert x s] x) (𝓝[insert x s ×ˢ insert x s] (x, x)) := by + rw [nhdsWithin_prod_eq] + exact Tendsto.prod_mk tendsto_id (tendsto_const_nhdsWithin (by simp)) + exact this (isLittleO_iff.1 h.hasStrictFDerivWithinAt hc) + theorem HasFPowerSeriesAt.hasFDerivAt (h : HasFPowerSeriesAt f p x) : HasFDerivAt f (continuousMultilinearCurryFin1 𝕜 E F (p 1)) x := h.hasStrictFDerivAt.hasFDerivAt +theorem HasFPowerSeriesWithinAt.differentiableWithinAt (h : HasFPowerSeriesWithinAt f p s x) : + DifferentiableWithinAt 𝕜 f (insert x s) x := + h.hasFDerivWithinAt.differentiableWithinAt + theorem HasFPowerSeriesAt.differentiableAt (h : HasFPowerSeriesAt f p x) : DifferentiableAt 𝕜 f x := h.hasFDerivAt.differentiableAt +theorem AnalyticWithinAt.differentiableWithinAt (h : AnalyticWithinAt 𝕜 f s x) : + DifferentiableWithinAt 𝕜 f (insert x s) x := by + obtain ⟨p, hp⟩ := h + exact hp.differentiableWithinAt + theorem AnalyticAt.differentiableAt : AnalyticAt 𝕜 f x → DifferentiableAt 𝕜 f x | ⟨_, hp⟩ => hp.differentiableAt theorem AnalyticAt.differentiableWithinAt (h : AnalyticAt 𝕜 f x) : DifferentiableWithinAt 𝕜 f s x := h.differentiableAt.differentiableWithinAt +theorem HasFPowerSeriesWithinAt.fderivWithin_eq + (h : HasFPowerSeriesWithinAt f p s x) (hu : UniqueDiffWithinAt 𝕜 (insert x s) x) : + fderivWithin 𝕜 f (insert x s) x = continuousMultilinearCurryFin1 𝕜 E F (p 1) := + h.hasFDerivWithinAt.fderivWithin hu + theorem HasFPowerSeriesAt.fderiv_eq (h : HasFPowerSeriesAt f p x) : fderiv 𝕜 f x = continuousMultilinearCurryFin1 𝕜 E F (p 1) := h.hasFDerivAt.fderiv @@ -64,29 +138,61 @@ theorem AnalyticAt.hasStrictFDerivAt (h : AnalyticAt 𝕜 f x) : rw [hp.fderiv_eq] exact hp.hasStrictFDerivAt +theorem HasFPowerSeriesWithinOnBall.differentiableOn [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) : + DifferentiableOn 𝕜 f (insert x s ∩ EMetric.ball x r) := by + intro y hy + have Z := (h.analyticWithinAt_of_mem hy).differentiableWithinAt + rcases eq_or_ne y x with rfl | hy + · exact Z.mono inter_subset_left + · apply (Z.mono (subset_insert _ _)).mono_of_mem + suffices s ∈ 𝓝[insert x s] y from nhdsWithin_mono _ inter_subset_left this + rw [nhdsWithin_insert_of_ne hy] + exact self_mem_nhdsWithin + theorem HasFPowerSeriesOnBall.differentiableOn [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) : DifferentiableOn 𝕜 f (EMetric.ball x r) := fun _ hy => (h.analyticAt_of_mem hy).differentiableWithinAt -theorem AnalyticOnNhd.differentiableOn (h : AnalyticOnNhd 𝕜 f s) : - DifferentiableOn 𝕜 f s := fun y hy => - (h y hy).differentiableWithinAt +theorem AnalyticOn.differentiableOn (h : AnalyticOn 𝕜 f s) : DifferentiableOn 𝕜 f s := + fun y hy ↦ (h y hy).differentiableWithinAt.mono (by simp) -@[deprecated (since := "2024-09-26")] -alias AnalyticOn.differentiableOn := AnalyticOnNhd.differentiableOn +theorem AnalyticOnNhd.differentiableOn (h : AnalyticOnNhd 𝕜 f s) : DifferentiableOn 𝕜 f s := + fun y hy ↦ (h y hy).differentiableWithinAt + +theorem HasFPowerSeriesWithinOnBall.hasFDerivWithinAt [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) + {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) (h'y : x + y ∈ insert x s) : + HasFDerivWithinAt f (continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1)) + (insert x s) (x + y) := by + rcases eq_or_ne y 0 with rfl | h''y + · convert (h.changeOrigin hy h'y).hasFPowerSeriesWithinAt.hasFDerivWithinAt + simp + · have Z := (h.changeOrigin hy h'y).hasFPowerSeriesWithinAt.hasFDerivWithinAt + apply (Z.mono (subset_insert _ _)).mono_of_mem + rw [nhdsWithin_insert_of_ne] + · exact self_mem_nhdsWithin + · simpa using h''y theorem HasFPowerSeriesOnBall.hasFDerivAt [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) : HasFDerivAt f (continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1)) (x + y) := (h.changeOrigin hy).hasFPowerSeriesAt.hasFDerivAt +theorem HasFPowerSeriesWithinOnBall.fderivWithin_eq [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) + {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) (h'y : x + y ∈ insert x s) (hu : UniqueDiffOn 𝕜 (insert x s)) : + fderivWithin 𝕜 f (insert x s) (x + y) = + continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1) := + (h.hasFDerivWithinAt hy h'y).fderivWithin (hu _ h'y) + theorem HasFPowerSeriesOnBall.fderiv_eq [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) : fderiv 𝕜 f (x + y) = continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1) := (h.hasFDerivAt hy).fderiv -/-- If a function has a power series on a ball, then so does its derivative. -/ -theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) : +protected theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] + (h : HasFPowerSeriesOnBall f p x r) : HasFPowerSeriesOnBall (fderiv 𝕜 f) p.derivSeries x r := by refine .congr (f := fun z ↦ continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin (z - x) 1)) ?_ fun z hz ↦ ?_ @@ -98,18 +204,42 @@ theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] (h : HasFPowerSeriesOnBal rw [← h.fderiv_eq, add_sub_cancel] simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz +/-- If a function has a power series within a set on a ball, then so does its derivative. -/ +protected theorem HasFPowerSeriesWithinOnBall.fderivWithin [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) (hu : UniqueDiffOn 𝕜 (insert x s)) : + HasFPowerSeriesWithinOnBall (fderivWithin 𝕜 f (insert x s)) p.derivSeries s x r := by + refine .congr' (f := fun z ↦ continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin (z - x) 1)) ?_ + (fun z hz ↦ ?_) + · refine continuousMultilinearCurryFin1 𝕜 E F + |>.toContinuousLinearEquiv.toContinuousLinearMap.comp_hasFPowerSeriesWithinOnBall ?_ + apply HasFPowerSeriesOnBall.hasFPowerSeriesWithinOnBall + simpa using ((p.hasFPowerSeriesOnBall_changeOrigin 1 + (h.r_pos.trans_le h.r_le)).mono h.r_pos h.r_le).comp_sub x + · dsimp only + rw [← h.fderivWithin_eq _ _ hu, add_sub_cancel] + · simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz.2 + · simpa using hz.1 + /-- If a function is analytic on a set `s`, so is its Fréchet derivative. -/ -theorem AnalyticOnNhd.fderiv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : - AnalyticOnNhd 𝕜 (fderiv 𝕜 f) s := by - intro y hy - rcases h y hy with ⟨p, r, hp⟩ +protected theorem AnalyticAt.fderiv [CompleteSpace F] (h : AnalyticAt 𝕜 f x) : + AnalyticAt 𝕜 (fderiv 𝕜 f) x := by + rcases h with ⟨p, r, hp⟩ exact hp.fderiv.analyticAt +/-- If a function is analytic on a set `s`, so is its Fréchet derivative. See also +`AnalyticOnNhd.fderiv_of_isOpen`, removing the completeness assumption but requiring the set +to be open. -/ +protected theorem AnalyticOnNhd.fderiv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : + AnalyticOnNhd 𝕜 (fderiv 𝕜 f) s := + fun y hy ↦ AnalyticAt.fderiv (h y hy) + @[deprecated (since := "2024-09-26")] alias AnalyticOn.fderiv := AnalyticOnNhd.fderiv -/-- If a function is analytic on a set `s`, so are its successive Fréchet derivative. -/ -theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (n : ℕ) : +/-- If a function is analytic on a set `s`, so are its successive Fréchet derivative. See also +`AnalyticOnNhd.iteratedFDeruv_of_isOpen`, removing the completeness assumption but requiring the set +to be open.-/ +protected theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (n : ℕ) : AnalyticOnNhd 𝕜 (iteratedFDeriv 𝕜 n f) s := by induction n with | zero => @@ -125,8 +255,21 @@ theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f @[deprecated (since := "2024-09-26")] alias AnalyticOn.iteratedFDeriv := AnalyticOnNhd.iteratedFDeriv +/-- If a function is analytic on a neighborhood of a set `s`, then it has a Taylor series given +by the sequence of its derivatives. Note that, if the function were just analytic on `s`, then +one would have to use instead the sequence of derivatives inside the set, as in +`AnalyticOn.hasFTaylorSeriesUpToOn`. -/ +lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] + (n : ℕ∞) (h : AnalyticOnNhd 𝕜 f s) : + HasFTaylorSeriesUpToOn n f (ftaylorSeries 𝕜 f) s := by + refine ⟨fun x _hx ↦ rfl, fun m _hm x hx ↦ ?_, fun m _hm x hx ↦ ?_⟩ + · apply HasFDerivAt.hasFDerivWithinAt + exact ((h.iteratedFDeriv m x hx).differentiableAt).hasFDerivAt + · apply (DifferentiableAt.continuousAt (𝕜 := 𝕜) ?_).continuousWithinAt + exact (h.iteratedFDeriv m x hx).differentiableAt + /-- An analytic function is infinitely differentiable. -/ -theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) {n : ℕ∞} : +protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s := let t := { x | AnalyticAt 𝕜 f x } suffices ContDiffOn 𝕜 n f t from this.mono h @@ -149,18 +292,139 @@ theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd exact hf.contDiffOn.contDiffAt hs -lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} +protected lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} (h : AnalyticWithinAt 𝕜 f s x) {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x := by rcases h.exists_analyticAt with ⟨g, fx, fg, hg⟩ exact hg.contDiffAt.contDiffWithinAt.congr (fg.mono (subset_insert _ _)) fx -lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E} +protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E} (h : AnalyticOn 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s := fun x m ↦ (h x m).contDiffWithinAt @[deprecated (since := "2024-09-26")] alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn +lemma AnalyticWithinAt.exists_hasFTaylorSeriesUpToOn [CompleteSpace F] + (n : ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : + ∃ u ∈ 𝓝[insert x s] x, ∃ (p : E → FormalMultilinearSeries 𝕜 E F), + HasFTaylorSeriesUpToOn n f p u ∧ ∀ i, AnalyticOn 𝕜 (fun x ↦ p x i) u := by + rcases h.exists_analyticAt with ⟨g, -, fg, hg⟩ + rcases hg.exists_mem_nhds_analyticOnNhd with ⟨v, vx, hv⟩ + refine ⟨insert x s ∩ v, inter_mem_nhdsWithin _ vx, ftaylorSeries 𝕜 g, ?_, fun i ↦ ?_⟩ + · suffices HasFTaylorSeriesUpToOn n g (ftaylorSeries 𝕜 g) (insert x s ∩ v) from + this.congr (fun y hy ↦ fg hy.1) + exact AnalyticOnNhd.hasFTaylorSeriesUpToOn _ (hv.mono Set.inter_subset_right) + · exact (hv.iteratedFDeriv i).analyticOn.mono Set.inter_subset_right + +/-- If a function has a power series `p` within a set of unique differentiability, inside a ball, +and is differentiable at a point, then the derivative series of `p` is summable at a point, with +sum the given differential. Note that this theorem does not require completeness of the space.-/ +theorem HasFPowerSeriesWithinOnBall.hasSum_derivSeries_of_hasFDerivWithinAt + (h : HasFPowerSeriesWithinOnBall f p s x r) + {f' : E →L[𝕜] F} + {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) (h'y : x + y ∈ insert x s) + (hf' : HasFDerivWithinAt f f' (insert x s) (x + y)) + (hu : UniqueDiffOn 𝕜 (insert x s)) : + HasSum (fun n ↦ p.derivSeries n (fun _ ↦ y)) f' := by + /- In the completion of the space, the derivative series is summable, and its sum is a derivative + of the function. Therefore, by uniqueness of derivatives, its sum is the image of `f'` under + the canonical embedding. As this is an embedding, it means that there was also convergence in + the original space, to `f'`. -/ + let F' := UniformSpace.Completion F + let a : F →L[𝕜] F' := UniformSpace.Completion.toComplL + let b : (E →L[𝕜] F) →ₗᵢ[𝕜] (E →L[𝕜] F') := UniformSpace.Completion.toComplₗᵢ.postcomp + rw [← b.embedding.hasSum_iff] + have : HasFPowerSeriesWithinOnBall (a ∘ f) (a.compFormalMultilinearSeries p) s x r := + a.comp_hasFPowerSeriesWithinOnBall h + have Z := (this.fderivWithin hu).hasSum h'y (by simpa [edist_eq_coe_nnnorm] using hy) + have : fderivWithin 𝕜 (a ∘ f) (insert x s) (x + y) = a ∘L f' := by + apply HasFDerivWithinAt.fderivWithin _ (hu _ h'y) + exact a.hasFDerivAt.comp_hasFDerivWithinAt (x + y) hf' + rw [this] at Z + convert Z with n + ext v + simp only [FormalMultilinearSeries.derivSeries, + ContinuousLinearMap.compFormalMultilinearSeries_apply, + FormalMultilinearSeries.changeOriginSeries, + ContinuousLinearMap.compContinuousMultilinearMap_coe, ContinuousLinearEquiv.coe_coe, + LinearIsometryEquiv.coe_coe, Function.comp_apply, ContinuousMultilinearMap.sum_apply, map_sum, + ContinuousLinearMap.coe_sum', Finset.sum_apply, + Matrix.zero_empty] + rfl + +/-- If a function is analytic within a set with unique differentials, then so is its derivative. +Note that this theorem does not require completeness of the space. -/ +protected theorem AnalyticOn.fderivWithin (h : AnalyticOn 𝕜 f s) (hu : UniqueDiffOn 𝕜 s) : + AnalyticOn 𝕜 (fderivWithin 𝕜 f s) s := by + intro x hx + rcases h x hx with ⟨p, r, hr⟩ + refine ⟨p.derivSeries, r, ?_⟩ + refine ⟨hr.r_le.trans p.radius_le_radius_derivSeries, hr.r_pos, fun {y} hy h'y ↦ ?_⟩ + apply hr.hasSum_derivSeries_of_hasFDerivWithinAt (by simpa [edist_eq_coe_nnnorm] using h'y) hy + rw [insert_eq_of_mem hx] at hy ⊢ + apply DifferentiableWithinAt.hasFDerivWithinAt + · exact h.differentiableOn _ hy + · rwa [insert_eq_of_mem hx] + +/-- If a function is analytic on a set `s`, so are its successive Fréchet derivative within this +set. Note that this theorem does not require completeness of the space. -/ +protected theorem AnalyticOn.iteratedFDerivWithin (h : AnalyticOn 𝕜 f s) + (hu : UniqueDiffOn 𝕜 s) (n : ℕ) : + AnalyticOn 𝕜 (iteratedFDerivWithin 𝕜 n f s) s := by + induction n with + | zero => + rw [iteratedFDerivWithin_zero_eq_comp] + exact ((continuousMultilinearCurryFin0 𝕜 E F).symm : F →L[𝕜] E[×0]→L[𝕜] F) + |>.comp_analyticOn h + | succ n IH => + rw [iteratedFDerivWithin_succ_eq_comp_left] + apply AnalyticOnNhd.comp_analyticOn _ (IH.fderivWithin hu) (mapsTo_univ _ _) + apply LinearIsometryEquiv.analyticOnNhd + +lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : ℕ∞} + (h : AnalyticOn 𝕜 f s) (hu : UniqueDiffOn 𝕜 s) : + HasFTaylorSeriesUpToOn n f (ftaylorSeriesWithin 𝕜 f s) s := by + refine ⟨fun x _hx ↦ rfl, fun m _hm x hx ↦ ?_, fun m _hm x hx ↦ ?_⟩ + · have := (h.iteratedFDerivWithin hu m x hx).differentiableWithinAt.hasFDerivWithinAt + rwa [insert_eq_of_mem hx] at this + · exact (h.iteratedFDerivWithin hu m x hx).continuousWithinAt + +lemma AnalyticOn.exists_hasFTaylorSeriesUpToOn + (h : AnalyticOn 𝕜 f s) (hu : UniqueDiffOn 𝕜 s) : + ∃ (p : E → FormalMultilinearSeries 𝕜 E F), + HasFTaylorSeriesUpToOn ⊤ f p s ∧ ∀ i, AnalyticOn 𝕜 (fun x ↦ p x i) s := + ⟨ftaylorSeriesWithin 𝕜 f s, h.hasFTaylorSeriesUpToOn hu, h.iteratedFDerivWithin hu⟩ + +theorem AnalyticOnNhd.fderiv_of_isOpen (h : AnalyticOnNhd 𝕜 f s) (hs : IsOpen s) : + AnalyticOnNhd 𝕜 (fderiv 𝕜 f) s := by + rw [← hs.analyticOn_iff_analyticOnNhd] at h ⊢ + exact (h.fderivWithin hs.uniqueDiffOn).congr (fun x hx ↦ (fderivWithin_of_isOpen hs hx).symm) + +theorem AnalyticOnNhd.iteratedFDeriv_of_isOpen (h : AnalyticOnNhd 𝕜 f s) (hs : IsOpen s) (n : ℕ) : + AnalyticOnNhd 𝕜 (iteratedFDeriv 𝕜 n f) s := by + rw [← hs.analyticOn_iff_analyticOnNhd] at h ⊢ + exact (h.iteratedFDerivWithin hs.uniqueDiffOn n).congr + (fun x hx ↦ (iteratedFDerivWithin_of_isOpen n hs hx).symm) + +/-- If a partial homeomorphism `f` is analytic at a point `a`, with invertible derivative, then +its inverse is analytic at `f a`. -/ +theorem PartialHomeomorph.analyticAt_symm' (f : PartialHomeomorph E F) {a : E} + {i : E ≃L[𝕜] F} (h0 : a ∈ f.source) (h : AnalyticAt 𝕜 f a) (h' : fderiv 𝕜 f a = i) : + AnalyticAt 𝕜 f.symm (f a) := by + rcases h with ⟨p, hp⟩ + have : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i := by simp [← h', hp.fderiv_eq] + exact (f.hasFPowerSeriesAt_symm h0 hp this).analyticAt + +/-- If a partial homeomorphism `f` is analytic at a point `f.symm a`, with invertible derivative, +then its inverse is analytic at `a`. -/ +theorem PartialHomeomorph.analyticAt_symm (f : PartialHomeomorph E F) {a : F} + {i : E ≃L[𝕜] F} (h0 : a ∈ f.target) (h : AnalyticAt 𝕜 f (f.symm a)) + (h' : fderiv 𝕜 f (f.symm a) = i) : + AnalyticAt 𝕜 f.symm a := by + have : a = f (f.symm a) := by simp [h0] + rw [this] + exact f.analyticAt_symm' (by simp [h0]) h h' + end fderiv section deriv @@ -180,14 +444,19 @@ protected theorem HasFPowerSeriesAt.deriv (h : HasFPowerSeriesAt f p x) : deriv f x = p 1 fun _ => 1 := h.hasDerivAt.deriv -/-- If a function is analytic on a set `s`, so is its derivative. -/ -theorem AnalyticOnNhd.deriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : +/-- If a function is analytic on a set `s` in a complete space, so is its derivative. -/ +protected theorem AnalyticOnNhd.deriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : AnalyticOnNhd 𝕜 (deriv f) s := (ContinuousLinearMap.apply 𝕜 F (1 : 𝕜)).comp_analyticOnNhd h.fderiv @[deprecated (since := "2024-09-26")] alias AnalyticOn.deriv := AnalyticOnNhd.deriv +/-- If a function is analytic on an open set `s`, so is its derivative. -/ +theorem AnalyticOnNhd.deriv_of_isOpen (h : AnalyticOnNhd 𝕜 f s) (hs : IsOpen s) : + AnalyticOnNhd 𝕜 (deriv f) s := + (ContinuousLinearMap.apply 𝕜 F (1 : 𝕜)).comp_analyticOnNhd (h.fderiv_of_isOpen hs) + /-- If a function is analytic on a set `s`, so are its successive derivatives. -/ theorem AnalyticOnNhd.iterated_deriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (n : ℕ) : AnalyticOnNhd 𝕜 (_root_.deriv^[n] f) s := by @@ -367,6 +636,30 @@ protected theorem hasFDerivAt [DecidableEq ι] : HasFDerivAt f (f.linearDeriv x) convert f.hasFiniteFPowerSeriesOnBall.hasFDerivAt (y := x) ENNReal.coe_lt_top rw [zero_add] +/-- Given `f` a multilinear map, then the derivative of `x ↦ f (g₁ x, ..., gₙ x)` at `x` applied +to a vector `v` is given by `∑ i, f (g₁ x, ..., g'ᵢ v, ..., gₙ x)`. Version inside a set. -/ +theorem _root_.HasFDerivWithinAt.multilinear_comp + [DecidableEq ι] {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] + {g : ∀ i, G → E i} {g' : ∀ i, G →L[𝕜] E i} {s : Set G} {x : G} + (hg : ∀ i, HasFDerivWithinAt (g i) (g' i) s x) : + HasFDerivWithinAt (fun x ↦ f (fun i ↦ g i x)) + ((∑ i : ι, (f.toContinuousLinearMap (fun j ↦ g j x) i) ∘L (g' i))) s x := by + convert (f.hasFDerivAt (fun j ↦ g j x)).comp_hasFDerivWithinAt x (hasFDerivWithinAt_pi.2 hg) + ext v + simp [linearDeriv] + +/-- Given `f` a multilinear map, then the derivative of `x ↦ f (g₁ x, ..., gₙ x)` at `x` applied +to a vector `v` is given by `∑ i, f (g₁ x, ..., g'ᵢ v, ..., gₙ x)`. -/ +theorem _root_.HasFDerivAt.multilinear_comp + [DecidableEq ι] {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] + {g : ∀ i, G → E i} {g' : ∀ i, G →L[𝕜] E i} {x : G} + (hg : ∀ i, HasFDerivAt (g i) (g' i) x) : + HasFDerivAt (fun x ↦ f (fun i ↦ g i x)) + ((∑ i : ι, (f.toContinuousLinearMap (fun j ↦ g j x) i) ∘L (g' i))) x := by + convert (f.hasFDerivAt (fun j ↦ g j x)).comp x (hasFDerivAt_pi.2 hg) + ext v + simp [linearDeriv] + /-- Technical lemma used in the proof of `hasFTaylorSeriesUpTo_iteratedFDeriv`, to compare sums over embedding of `Fin k` and `Fin (k + 1)`. -/ private lemma _root_.Equiv.succ_embeddingFinSucc_fst_symm_apply {ι : Type*} [DecidableEq ι] @@ -521,3 +814,73 @@ theorem hasSum_iteratedFDeriv [CharZero 𝕜] {y : E} (hy : y ∈ EMetric.ball 0 mul_inv_cancel₀ <| cast_ne_zero.mpr n.factorial_ne_zero, one_smul] end HasFPowerSeriesOnBall + +/-! +### Derivative of a linear map into multilinear maps +-/ + +namespace ContinuousLinearMap + +variable {ι : Type*} {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, NormedSpace 𝕜 (G i)] + [Fintype ι] {H : Type*} [NormedAddCommGroup H] + [NormedSpace 𝕜 H] + +theorem hasFDerivAt_uncurry_of_multilinear [DecidableEq ι] + (f : E →L[𝕜] ContinuousMultilinearMap 𝕜 G F) (v : E × Π i, G i) : + HasFDerivAt (fun (p : E × Π i, G i) ↦ f p.1 p.2) + ((f.flipMultilinear v.2) ∘L (.fst _ _ _) + + ∑ i : ι, ((f v.1).toContinuousLinearMap v.2 i) ∘L (.proj _) ∘L (.snd _ _ _)) v := by + convert HasFDerivAt.multilinear_comp (f.continuousMultilinearMapOption) + (g := fun (_ : Option ι) p ↦ p) (g' := fun _ ↦ ContinuousLinearMap.id _ _) (x := v) + (fun _ ↦ hasFDerivAt_id _) + have I : f.continuousMultilinearMapOption.toContinuousLinearMap (fun _ ↦ v) none = + (f.flipMultilinear v.2) ∘L (.fst _ _ _) := by + simp [ContinuousMultilinearMap.toContinuousLinearMap, continuousMultilinearMapOption] + apply ContinuousLinearMap.ext (fun w ↦ ?_) + simp + have J : ∀ (i : ι), f.continuousMultilinearMapOption.toContinuousLinearMap (fun _ ↦ v) (some i) + = ((f v.1).toContinuousLinearMap v.2 i) ∘L (.proj _) ∘L (.snd _ _ _) := by + intro i + apply ContinuousLinearMap.ext (fun w ↦ ?_) + simp only [ContinuousMultilinearMap.toContinuousLinearMap, continuousMultilinearMapOption, + coe_mk', MultilinearMap.toLinearMap_apply, ContinuousMultilinearMap.coe_coe, + MultilinearMap.coe_mkContinuous, MultilinearMap.coe_mk, ne_eq, reduceCtorEq, + not_false_eq_true, Function.update_noteq, coe_comp', coe_snd', Function.comp_apply, + proj_apply] + congr + ext j + rcases eq_or_ne j i with rfl | hij + · simp + · simp [hij] + simp [I, J] + +/-- Given `f` a linear map into multilinear maps, then the derivative +of `x ↦ f (a x) (b₁ x, ..., bₙ x)` at `x` applied to a vector `v` is given by +`f (a' v) (b₁ x, ...., bₙ x) + ∑ i, f a (b₁ x, ..., b'ᵢ v, ..., bₙ x)`. Version inside a set. -/ +theorem _root_.HasFDerivWithinAt.linear_multilinear_comp + [DecidableEq ι] {a : H → E} {a' : H →L[𝕜] E} + {b : ∀ i, H → G i} {b' : ∀ i, H →L[𝕜] G i} {s : Set H} {x : H} + (ha : HasFDerivWithinAt a a' s x) (hb : ∀ i, HasFDerivWithinAt (b i) (b' i) s x) + (f : E →L[𝕜] ContinuousMultilinearMap 𝕜 G F) : + HasFDerivWithinAt (fun y ↦ f (a y) (fun i ↦ b i y)) + ((f.flipMultilinear (fun i ↦ b i x)) ∘L a' + + ∑ i, ((f (a x)).toContinuousLinearMap (fun j ↦ b j x) i) ∘L (b' i)) s x := by + convert (hasFDerivAt_uncurry_of_multilinear f (a x, fun i ↦ b i x)).comp_hasFDerivWithinAt x + (ha.prod (hasFDerivWithinAt_pi.mpr hb)) + ext v + simp + +/-- Given `f` a linear map into multilinear maps, then the derivative +of `x ↦ f (a x) (b₁ x, ..., bₙ x)` at `x` applied to a vector `v` is given by +`f (a' v) (b₁ x, ...., bₙ x) + ∑ i, f a (b₁ x, ..., b'ᵢ v, ..., bₙ x)`. -/ +theorem _root_.HasFDerivAt.linear_multilinear_comp [DecidableEq ι] {a : H → E} {a' : H →L[𝕜] E} + {b : ∀ i, H → G i} {b' : ∀ i, H →L[𝕜] G i} {x : H} + (ha : HasFDerivAt a a' x) (hb : ∀ i, HasFDerivAt (b i) (b' i) x) + (f : E →L[𝕜] ContinuousMultilinearMap 𝕜 G F) : + HasFDerivAt (fun y ↦ f (a y) (fun i ↦ b i y)) + ((f.flipMultilinear (fun i ↦ b i x)) ∘L a' + + ∑ i, ((f (a x)).toContinuousLinearMap (fun j ↦ b j x) i) ∘L (b' i)) x := by + simp_rw [← hasFDerivWithinAt_univ] at ha hb ⊢ + exact HasFDerivWithinAt.linear_multilinear_comp ha hb f + +end ContinuousLinearMap diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 578fd81859e6d..77c5f4cf43a72 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -205,6 +205,8 @@ theorem nnnorm_map (x : E) : ‖f x‖₊ = ‖x‖₊ := protected theorem isometry : Isometry f := AddMonoidHomClass.isometry_of_norm f.toLinearMap (norm_map _) +protected lemma embedding (f : F →ₛₗᵢ[σ₁₂] E₂) : Embedding f := f.isometry.embedding + -- Should be `@[simp]` but it doesn't fire due to `lean4#3107`. theorem isComplete_image_iff [SemilinearIsometryClass 𝓕 σ₁₂ E E₂] (f : 𝓕) {s : Set E} : IsComplete (f '' s) ↔ IsComplete s := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index 85edc3f7a843e..253805834044d 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -165,6 +165,15 @@ theorem norm_toContinuousLinearMap_comp [RingHomIsometric σ₁₂] (f : F → opNorm_ext (f.toContinuousLinearMap.comp g) g fun x => by simp only [norm_map, coe_toContinuousLinearMap, coe_comp', Function.comp_apply] +/-- Composing on the left with a linear isometry gives a linear isometry between spaces of +continuous linear maps. -/ +def postcomp [RingHomIsometric σ₁₂] [RingHomIsometric σ₁₃] (a : F →ₛₗᵢ[σ₂₃] G) : + (E →SL[σ₁₂] F) →ₛₗᵢ[σ₂₃] (E →SL[σ₁₃] G) where + toFun f := a.toContinuousLinearMap.comp f + map_add' f g := by simp + map_smul' c f := by simp + norm_map' f := by simp [a.norm_toContinuousLinearMap_comp] + end LinearIsometry end From 3ad544d9d2e7bbd808fb00b0591cb8513764bc24 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 4 Oct 2024 07:26:49 +0000 Subject: [PATCH 174/174] ci(nightly_detect_failure): add actionable message in case of failure (#17329) Also, remove the logic that detects whether a message has already been posted, since we now post the sha which is always different. --- .github/workflows/nightly_detect_failure.yml | 30 +++++++------------- 1 file changed, 10 insertions(+), 20 deletions(-) diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index 140f115408e95..25835698e99ab 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -23,6 +23,7 @@ jobs: topic: 'Mathlib status updates' content: | ❌ The latest CI for Mathlib's branch#nightly-testing has [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.event.workflow_run.id }}) ([${{ github.sha }}](https://github.com/${{ github.repository }}/commit/${{ github.sha }})). + You can `git fetch; git checkout nightly-testing` and push a fix. handle_success: if: ${{ github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.head_branch == 'nightly-testing' }} @@ -211,25 +212,14 @@ jobs: payload = f"🛠️: it looks like it's time to create a new bump/nightly-{current_version} branch from nightly-testing (specifically {sha}), and then PR that to {bump_branch}. " payload += "To do so semi-automatically, run the following script from mathlib root:\n\n" payload += f"```bash\n./scripts/create-adaptation-pr.sh --bumpversion={bump_branch_suffix} --nightlydate={current_version} --nightlysha={sha}\n```\n" - # Only post if the message is different - # We compare the first 160 characters, since that includes the date and bump version - if not messages or messages[0]['content'][:160] != payload[:160]: - # Log messages, because the bot seems to repeat itself... - if messages: - print("###### Last message:") - print(messages[0]['content']) - print("###### Current message:") - print(payload) - else: - print('The strings match!') - # Post the reminder message - request = { - 'type': 'stream', - 'to': 'nightly-testing', - 'topic': 'Mathlib bump branch reminders', - 'content': payload - } - result = client.send_message(request) - print(result) + # Post the reminder message + request = { + 'type': 'stream', + 'to': 'nightly-testing', + 'topic': 'Mathlib bump branch reminders', + 'content': payload + } + result = client.send_message(request) + print(result) else: print('No action needed.')